From a42795cc1c2a8ed3efa9960af920ff7b16d928f0 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Thu, 1 Sep 2016 17:52:44 +0200 Subject: Introducing a new EConstr.t type to perform the nf_evar operation on demand. --- dev/db | 1 + dev/top_printers.ml | 1 + engine/eConstr.ml | 614 ++++++++++++++++++++++++++++++++++++++++++++++++++++ engine/eConstr.mli | 183 ++++++++++++++++ engine/engine.mllib | 1 + kernel/constr.ml | 6 + kernel/constr.mli | 10 + 7 files changed, 816 insertions(+) create mode 100644 engine/eConstr.ml create mode 100644 engine/eConstr.mli diff --git a/dev/db b/dev/db index d3503ef43..432baf8a6 100644 --- a/dev/db +++ b/dev/db @@ -28,6 +28,7 @@ install_printer Top_printers.pppattern install_printer Top_printers.ppglob_constr install_printer Top_printers.ppconstr +install_printer Top_printers.ppeconstr install_printer Top_printers.ppuni install_printer Top_printers.ppuniverses install_printer Top_printers.ppconstraints diff --git a/dev/top_printers.ml b/dev/top_printers.ml index b552d9994..b7736f498 100644 --- a/dev/top_printers.ml +++ b/dev/top_printers.ml @@ -71,6 +71,7 @@ let ppwf_paths x = pp (Rtree.pp_tree pprecarg x) let rawdebug = ref false let ppevar evk = pp (str (Evd.string_of_existential evk)) let ppconstr x = pp (Termops.print_constr x) +let ppeconstr x = pp (Termops.print_constr (EConstr.Unsafe.to_constr x)) let ppconstr_expr x = pp (Ppconstr.pr_constr_expr x) let ppconstrdb x = pp(Flags.with_option rawdebug Termops.print_constr x) let ppterm = ppconstr diff --git a/engine/eConstr.ml b/engine/eConstr.ml new file mode 100644 index 000000000..0a5f1ba68 --- /dev/null +++ b/engine/eConstr.ml @@ -0,0 +1,614 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* t -> (t, t) Constr.kind_of_term +val kind_upto : Evd.evar_map -> constr -> (constr, types) Constr.kind_of_term +val of_kind : (t, t) Constr.kind_of_term -> t +val of_constr : Constr.t -> t +val unsafe_to_constr : t -> Constr.t +val unsafe_eq : (t, Constr.t) eq +end = +struct + +type t = Constr.t + +let safe_evar_value sigma ev = + try Some (Evd.existential_value sigma ev) + with NotInstantiatedEvar | Not_found -> None + +let rec kind sigma c = + let c0 = Constr.kind c in + match c0 with + | Evar (evk, args) -> + (match safe_evar_value sigma (evk, args) with + Some c -> kind sigma c + | None -> c0) + | Sort (Type u) -> + let u' = Evd.normalize_universe sigma u in + if u' == u then c0 else Sort (Type u') + | Const (c', u) when not (Univ.Instance.is_empty u) -> + let u' = Evd.normalize_universe_instance sigma u in + if u' == u then c0 else Const (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 c0 else Ind (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 c0 else Construct (co, u') + | App (c, args) when isEvar c -> + (** Enforce smart constructor invariant on applications *) + let ev = destEvar c in + begin match safe_evar_value sigma ev with + | None -> c0 + | Some c -> kind sigma (mkApp (c, args)) + end + | _ -> c0 + +let kind_upto = kind +let of_kind = Constr.of_kind +let of_constr c = c +let unsafe_to_constr c = c +let unsafe_eq = Refl + +end + +include API + +type types = t +type constr = t +type existential = t pexistential +type fixpoint = (t, t) pfixpoint +type cofixpoint = (t, t) pcofixpoint + +let mkProp = of_kind (Sort Sorts.prop) +let mkSet = of_kind (Sort Sorts.set) +let mkType u = of_kind (Sort (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 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 mkIndU pi = of_kind (Ind pi) +let mkConstructU pc = of_kind (Construct pc) +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 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 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 local_assum (na, t) = + let open Rel.Declaration in + LocalAssum (na, unsafe_to_constr t) + +let local_def (na, b, t) = + let open Rel.Declaration in + LocalDef (na, unsafe_to_constr b, unsafe_to_constr t) + +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 rec lamdec_rec l c = + match kind sigma c with + | Lambda (x,t,c) -> lamdec_rec (Context.Rel.add (local_assum (x,t)) l) c + | LetIn (x,b,t,c) -> lamdec_rec (Context.Rel.add (local_def (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 = + if n < 0 then + error "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 (local_assum (x,t)) l) (n-1) c + | LetIn (x,b,t,c) -> lamdec_rec (Context.Rel.add (local_def (x,b,t)) l) n c + | Cast (c,_,_) -> lamdec_rec l n c + | c -> error "decompose_lam_n_assum: not enough abstractions" + in + lamdec_rec Context.Rel.empty n c + +let decompose_lam_n_decls sigma n = + if n < 0 then + error "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 (local_assum (x,t)) l) (n-1) c + | LetIn (x,b,t,c) -> lamdec_rec (Context.Rel.add (local_def (x,b,t)) l) (n-1) c + | Cast (c,_,_) -> lamdec_rec l n c + | c -> error "decompose_lam_n_decls: not enough abstractions" + in + lamdec_rec Context.Rel.empty n + +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 rec proddec_rec l c = + match kind sigma c with + | Prod (x,t,c) -> proddec_rec (Context.Rel.add (local_assum (x,t)) l) c + | LetIn (x,b,t,c) -> proddec_rec (Context.Rel.add (local_def (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 = + if n < 0 then + error "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 (local_assum (x,t)) l) (n-1) c + | LetIn (x,b,t,c) -> prodec_rec (Context.Rel.add (local_def (x,b,t)) l) (n-1) c + | Cast (c,_,_) -> prodec_rec l n c + | c -> error "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 = 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 (local_assum (na, t)) n) c + | Lambda (na,t,c) -> f n t; f (g (local_assum (na, t)) n) c + | LetIn (na,b,t,c) -> f n b; f n t; f (g (local_def (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 (local_assum (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 (local_assum (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 to_constr sigma t = + let rec map c = Constr.map map (Constr.of_kind (kind_upto sigma c)) in + map (unsafe_to_constr t) + +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) + +(** 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 + +module Vars = +struct +exception LocalOccur +let to_constr = unsafe_to_constr + +(** 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 (List.map to_constr subst) n (to_constr c)) +let substl subst c = of_constr (Vars.substl (List.map to_constr subst) (to_constr c)) +let subst1 c r = of_constr (Vars.subst1 (to_constr c) (to_constr r)) + +let replace_vars subst c = + let map (id, c) = (id, to_constr c) in + of_constr (Vars.replace_vars (List.map map 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)) + +(** 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 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 + +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, of_constr t, c) + | LocalDef (na,b,t) -> mkLetIn (na, of_constr b, of_constr t, c) + +let mkLambda_or_LetIn decl c = + let open Context.Rel.Declaration in + match decl with + | LocalAssum (na,t) -> mkLambda (na, of_constr t, c) + | LocalDef (na,b,t) -> mkLetIn (na, of_constr b, of_constr 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 + +module Unsafe = +struct +let to_constr = unsafe_to_constr +let eq = unsafe_eq +end diff --git a/engine/eConstr.mli b/engine/eConstr.mli new file mode 100644 index 000000000..03e2d4ffc --- /dev/null +++ b/engine/eConstr.mli @@ -0,0 +1,183 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* t -> (t, t) Constr.kind_of_term +(** Same as {!Constr.kind} except that it expands evars and normalizes + universes on the fly. *) + +val to_constr : Evd.evar_map -> t -> Constr.t +(** Returns the evar-normal form of the argument. See {!Evarutil.nf_evar}. *) + +(** {5 Constructors} *) + +val of_kind : (t, 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 : pconstant -> t +val mkProj : (projection * t) -> t +(* val mkInd : inductive -> t *) +val mkIndU : pinductive -> t +(* val mkConstruct : constructor -> t *) +val mkConstructU : pconstructor -> t +(* val mkConstructUi : pinductive * int -> t *) +val mkCase : case_info * t * t * t array -> t +val mkFix : (t, t) pfixpoint -> t +val mkCoFix : (t, t) pcofixpoint -> t + +val applist : t * t list -> t + +val mkProd_or_LetIn : Rel.Declaration.t -> t -> t +val mkLambda_or_LetIn : Rel.Declaration.t -> t -> t +val it_mkProd_or_LetIn : t -> Rel.t -> t +val it_mkLambda_or_LetIn : t -> Rel.t -> t + +(** {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 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 -> Sorts.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 puniverses +val destEvar : Evd.evar_map -> t -> t pexistential +val destInd : Evd.evar_map -> t -> inductive puniverses +val destConstruct : Evd.evar_map -> t -> constructor puniverses +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 -> Context.Rel.t * t +val decompose_lam_n_assum : Evd.evar_map -> int -> t -> Context.Rel.t * t +val decompose_lam_n_decls : Evd.evar_map -> int -> t -> Context.Rel.t * t + +val decompose_prod : Evd.evar_map -> t -> (Name.t * t) list * t +val decompose_prod_assum : Evd.evar_map -> t -> Context.Rel.t * t +val decompose_prod_n_assum : Evd.evar_map -> int -> t -> Context.Rel.t * t + +val existential_type : Evd.evar_map -> existential -> types + +(** {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 + +(** {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.t -> 'a -> 'a) -> ('a -> t -> unit) -> 'a -> t -> unit +val fold : Evd.evar_map -> ('a -> t -> 'a) -> 'a -> t -> 'a + +(** {6 Substitutions} *) + +module Vars : +sig +val lift : int -> t -> t +val liftn : int -> int -> t -> t +val substnl : t list -> int -> t -> t +val substl : t list -> t -> t +val subst1 : t -> t -> t + +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 +end + +(** {5 Unsafe operations} *) + +module Unsafe : +sig + val to_constr : t -> Constr.t + (** Physical identity. Does not care for defined evars. *) + + val eq : (t, Constr.t) eq + (** Use for transparent cast between types. *) +end diff --git a/engine/engine.mllib b/engine/engine.mllib index 53cbbd73e..c78f37a85 100644 --- a/engine/engine.mllib +++ b/engine/engine.mllib @@ -4,6 +4,7 @@ Universes UState Evd Sigma +EConstr Termops Proofview_monad Evarutil diff --git a/kernel/constr.ml b/kernel/constr.ml index ce20751ab..0ae3fb474 100644 --- a/kernel/constr.ml +++ b/kernel/constr.ml @@ -235,6 +235,12 @@ let mkVar id = Var id let kind c = c +(* The other way around. We treat specifically smart constructors *) +let of_kind = function +| App (f, a) -> mkApp (f, a) +| Cast (c, knd, t) -> mkCast (c, knd, t) +| k -> k + (****************************************************************************) (* Functions to recur through subterms *) (****************************************************************************) diff --git a/kernel/constr.mli b/kernel/constr.mli index 7095dbe6f..c8be89fe2 100644 --- a/kernel/constr.mli +++ b/kernel/constr.mli @@ -223,6 +223,7 @@ type ('constr, 'types) kind_of_term = term *) val kind : constr -> (constr, types) kind_of_term +val of_kind : (constr, types) kind_of_term -> constr (** [equal a b] is true if [a] equals [b] modulo alpha, casts, and application grouping *) @@ -310,6 +311,15 @@ val compare_head_gen : (bool -> Univ.Instance.t -> Univ.Instance.t -> bool) -> (constr -> constr -> bool) -> constr -> constr -> bool +val compare_head_gen_leq_with : + (constr -> (constr,types) kind_of_term) -> + (constr -> (constr,types) kind_of_term) -> + (bool -> Univ.Instance.t -> Univ.Instance.t -> bool) -> + (Sorts.t -> Sorts.t -> bool) -> + (constr -> constr -> bool) -> + (constr -> constr -> bool) -> + constr -> constr -> bool + (** [compare_head_gen_with k1 k2 u s f c1 c2] compares [c1] and [c2] like [compare_head_gen u s f c1 c2], except that [k1] (resp. [k2]) is used,rather than {!kind}, to expose the immediate subterms of -- cgit v1.2.3 From 5143129baac805d3a49ac3ee9f3344c7a447634f Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sun, 30 Oct 2016 17:53:07 +0100 Subject: Termops API using EConstr. --- engine/evarutil.ml | 4 +- engine/termops.ml | 303 +++++++++++++------------ engine/termops.mli | 94 ++++---- ltac/extratactics.ml4 | 4 +- ltac/rewrite.ml | 11 +- ltac/tauto.ml | 19 +- plugins/btauto/refl_btauto.ml | 19 +- plugins/cc/cctac.ml | 10 +- plugins/decl_mode/decl_interp.ml | 2 +- plugins/decl_mode/decl_proof_instr.ml | 32 +-- plugins/extraction/extraction.ml | 6 +- plugins/firstorder/formula.ml | 12 +- plugins/firstorder/rules.ml | 4 +- plugins/firstorder/unify.ml | 13 +- plugins/fourier/fourierR.ml | 3 +- plugins/funind/functional_principles_proofs.ml | 22 +- plugins/funind/functional_principles_types.ml | 18 +- plugins/funind/indfun.ml | 9 +- plugins/funind/invfun.ml | 8 +- plugins/funind/merge.ml | 4 +- plugins/funind/recdef.ml | 12 +- plugins/micromega/coq_micromega.ml | 5 +- plugins/quote/quote.ml | 8 +- plugins/rtauto/refl_tauto.ml | 4 +- plugins/ssrmatching/ssrmatching.ml4 | 2 +- pretyping/cases.ml | 89 ++++---- pretyping/coercion.ml | 2 +- pretyping/constr_matching.ml | 20 +- pretyping/detyping.ml | 32 +-- pretyping/evarconv.ml | 23 +- pretyping/evarsolve.ml | 71 +++--- pretyping/evarsolve.mli | 4 +- pretyping/pretyping.ml | 4 +- pretyping/recordops.ml | 6 +- pretyping/recordops.mli | 2 +- pretyping/reductionops.ml | 41 ++-- pretyping/reductionops.mli | 2 +- pretyping/retyping.ml | 17 +- pretyping/tacred.ml | 14 +- pretyping/unification.ml | 46 ++-- printing/printer.ml | 3 +- proofs/clenv.ml | 14 +- proofs/clenvtac.ml | 6 +- proofs/evar_refiner.ml | 2 +- proofs/logic.ml | 56 ++--- proofs/logic.mli | 2 +- proofs/redexpr.ml | 2 +- tactics/auto.ml | 18 +- tactics/class_tactics.ml | 29 +-- tactics/contradiction.ml | 10 +- tactics/eauto.ml | 17 +- tactics/elim.ml | 11 +- tactics/equality.ml | 41 ++-- tactics/hints.ml | 40 ++-- tactics/hints.mli | 2 +- tactics/hipattern.ml | 96 ++++---- tactics/hipattern.mli | 4 +- tactics/inv.ml | 17 +- tactics/leminv.ml | 7 +- tactics/tacticals.ml | 2 +- tactics/tactics.ml | 125 +++++----- tactics/term_dnet.ml | 2 +- toplevel/assumptions.ml | 25 +- toplevel/class.ml | 2 +- toplevel/command.ml | 8 +- toplevel/command.mli | 2 +- toplevel/himsg.ml | 2 +- toplevel/indschemes.ml | 4 +- toplevel/obligations.ml | 6 +- toplevel/record.ml | 2 +- toplevel/search.ml | 4 +- toplevel/vernacentries.ml | 2 +- 72 files changed, 825 insertions(+), 739 deletions(-) diff --git a/engine/evarutil.ml b/engine/evarutil.ml index 13444cb37..05f98a41f 100644 --- a/engine/evarutil.ml +++ b/engine/evarutil.ml @@ -540,12 +540,12 @@ 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 diff --git a/engine/termops.ml b/engine/termops.ml index 35917b368..356312e2f 100644 --- a/engine/termops.ml +++ b/engine/termops.ml @@ -236,35 +236,35 @@ 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 - (mkApp (f,fst (Array.chop (Array.length args - 1) args))) - | _ -> c + | App (f,args) when EConstr.isEvar sigma (Array.last args) -> + drop_extra_implicit_args sigma + (EConstr.mkApp (f,fst (Array.chop (Array.length args - 1) args))) + | _ -> EConstr.Unsafe.to_constr c (* Get the last arg of an application *) -let last_arg c = match kind_of_term c with - | App (f,cl) -> Array.last cl +let last_arg sigma c = match EConstr.kind sigma c with + | App (f,cl) -> EConstr.Unsafe.to_constr (Array.last cl) | _ -> anomaly (Pp.str "last_arg") (* Get the last arg of an application *) -let decompose_app_vect c = - match kind_of_term c with - | App (f,cl) -> (f, cl) - | _ -> (c,[||]) +let decompose_app_vect sigma c = + match EConstr.kind sigma c with + | App (f,cl) -> (EConstr.Unsafe.to_constr f, Array.map EConstr.Unsafe.to_constr cl) + | _ -> (EConstr.Unsafe.to_constr c,[||]) let adjust_app_list_size f1 l1 f2 l2 = let len1 = List.length l1 and len2 = List.length l2 in @@ -400,9 +400,11 @@ 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 inj c = EConstr.Unsafe.to_constr c in + 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) -> @@ -411,16 +413,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, inj 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, inj 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, inj b, inj 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 @@ -441,7 +443,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, inj 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 @@ -449,7 +451,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, inj 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 @@ -462,30 +464,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 @@ -520,29 +523,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 @@ -550,55 +553,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 (EConstr.of_constr typ) | LocalDef (_, body, typ) -> - occur_var env hyp typ || - occur_var env hyp body + occur_var env sigma hyp (EConstr.of_constr typ) || + occur_var env sigma hyp (EConstr.of_constr 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 = @@ -608,54 +611,54 @@ 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 dependent_main noevar univs sigma m t = let eqc x y = - if univs then not (Option.is_empty (Universes.eq_constr_universes x y)) - else eq_constr_nounivs x y + if univs then not (Option.is_empty (EConstr.eq_constr_universes sigma x y)) + else EConstr.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))); + deprec m (EConstr.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 && EConstr.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 -> EConstr.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 (EConstr.of_constr t) + | LocalDef (_, body, t) -> dependent sigma a (EConstr.of_constr body) || dependent sigma a (EConstr.of_constr t) -let count_occurrences m t = +let count_occurrences sigma m t = 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))); + countrec m (EConstr.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 EConstr.isMeta sigma c -> () | _, Evar _ -> () - | _ -> iter_constr_with_binders (lift 1) countrec m t + | _ -> EConstr.iter_with_binders sigma (EConstr.Vars.lift 1) countrec m t in countrec m t; !n @@ -663,7 +666,7 @@ let count_occurrences m t = (* Synonymous *) let occur_term = dependent -let pop t = lift (-1) t +let pop t = EConstr.Unsafe.to_constr (EConstr.Vars.lift (-1) t) (***************************) (* bindings functions *) @@ -678,45 +681,45 @@ let rec subst_meta bl c = | Meta i -> (try Int.List.assoc i bl with Not_found -> c) | _ -> map_constr (subst_meta bl) c -let rec strip_outer_cast c = match kind_of_term c with - | Cast (c,_,_) -> strip_outer_cast c - | _ -> c +let rec strip_outer_cast sigma c = match EConstr.kind sigma c with + | Cast (c,_,_) -> strip_outer_cast sigma c + | _ -> EConstr.Unsafe.to_constr c (* flattens application lists throwing casts in-between *) -let collapse_appl c = match kind_of_term c with +let collapse_appl sigma c = match EConstr.kind sigma c with | App (f,cl) -> let rec collapse_rec f cl2 = - match kind_of_term (strip_outer_cast f) with + match EConstr.kind sigma (EConstr.of_constr (strip_outer_cast sigma f)) with | App (g,cl1) -> collapse_rec g (Array.append cl1 cl2) - | _ -> mkApp (f,cl2) + | _ -> EConstr.mkApp (f,cl2) in - collapse_rec f cl - | _ -> c + EConstr.Unsafe.to_constr (collapse_rec f cl) + | _ -> EConstr.Unsafe.to_constr 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 c' = EConstr.of_constr (collapse_appl sigma c) and t' = EConstr.of_constr (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 (mkRel k, Array.sub cl2 l1 (l2 - l1))) + && eq_fun sigma c' (EConstr.mkApp (f2, Array.sub cl2 0 l1)) then + Some (EConstr.mkApp (EConstr.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 c' = EConstr.of_constr (collapse_appl sigma c) and t' = EConstr.of_constr (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' (EConstr.mkApp (f2, Array.sub cl2 0 l1)) then + Some (EConstr.mkApp ((EConstr.Vars.lift k by_c), Array.sub cl2 l1 (l2 - l1))) else None | _ -> None @@ -725,35 +728,35 @@ 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 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 EConstr.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, EConstr.Vars.lift 1 c)) substrec kc t in - substrec (1,c) t + EConstr.Unsafe.to_constr (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 + EConstr.Unsafe.to_constr (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 = @@ -804,13 +807,13 @@ 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 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 +let is_template_polymorphic env sigma f = + match EConstr.kind sigma f with | Ind ind -> Environ.template_polymorphic_pind ind env | Const c -> Environ.template_polymorphic_pconstant c env | _ -> false @@ -882,45 +885,46 @@ let filtering env cv_pb c1 c2 = 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 - | 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 [] +let decompose_prod_letin sigma c = + let inj c = EConstr.Unsafe.to_constr c in + let rec prodec_rec i l sigma c = match EConstr.kind sigma c with + | Prod (n,t,c) -> prodec_rec (succ i) (RelDecl.LocalAssum (n, inj t)::l) sigma c + | LetIn (n,d,t,c) -> prodec_rec (succ i) (RelDecl.LocalDef (n, inj d, inj t)::l) sigma c + | Cast (c,_,_) -> prodec_rec i l sigma c + | _ -> i,l, inj c in + prodec_rec 0 [] sigma 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 : Context.Rel.t * constr = + 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 @@ -1031,22 +1035,33 @@ 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.Unsafe.to_constr 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 (EConstr.of_constr 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 (EConstr.of_constr t)) + (global_vars_set env sigma (EConstr.of_constr 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,[]) diff --git a/engine/termops.mli b/engine/termops.mli index 78826f79a..5d53ce09e 100644 --- a/engine/termops.mli +++ b/engine/termops.mli @@ -69,8 +69,9 @@ val map_constr_with_binders_left_to_right : ('a -> constr -> constr) -> 'a -> constr -> constr val map_constr_with_full_binders : + Evd.evar_map -> (Context.Rel.Declaration.t -> 'a -> 'a) -> - ('a -> constr -> constr) -> 'a -> constr -> constr + ('a -> EConstr.t -> EConstr.t) -> 'a -> EConstr.t -> EConstr.t (** [fold_constr_with_binders g f n acc c] folds [f n] on the immediate subterms of [c] starting from [acc] and proceeding from left to @@ -80,11 +81,12 @@ val map_constr_with_full_binders : each binder traversal; it is not recursive *) val fold_constr_with_binders : - ('a -> 'a) -> ('a -> 'b -> constr -> 'b) -> 'a -> 'b -> constr -> 'b + Evd.evar_map -> ('a -> 'a) -> + ('a -> 'b -> EConstr.t -> 'b) -> 'a -> 'b -> EConstr.t -> 'b val fold_constr_with_full_binders : - (Context.Rel.Declaration.t -> 'a -> 'a) -> ('a -> 'b -> constr -> 'b) -> - 'a -> 'b -> constr -> 'b + Evd.evar_map -> (Context.Rel.Declaration.t -> 'a -> 'a) -> + ('a -> 'b -> EConstr.t -> 'b) -> 'a -> 'b -> EConstr.t -> 'b val iter_constr_with_full_binders : (Context.Rel.Declaration.t -> 'a -> 'a) -> ('a -> constr -> unit) -> 'a -> @@ -92,39 +94,39 @@ val iter_constr_with_full_binders : (**********************************************************************) -val strip_head_cast : constr -> constr -val drop_extra_implicit_args : constr -> constr +val strip_head_cast : Evd.evar_map -> EConstr.t -> EConstr.t +val drop_extra_implicit_args : Evd.evar_map -> EConstr.t -> 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 -> EConstr.t -> bool +val occur_existential : Evd.evar_map -> EConstr.t -> bool +val occur_meta_or_existential : Evd.evar_map -> EConstr.t -> bool +val occur_evar : Evd.evar_map -> existential_key -> EConstr.t -> bool +val occur_var : env -> Evd.evar_map -> Id.t -> EConstr.t -> bool val occur_var_in_decl : - env -> + env -> Evd.evar_map -> Id.t -> Context.Named.Declaration.t -> 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 -> EConstr.t -> bool -val free_rels : constr -> Int.Set.t +val free_rels : Evd.evar_map -> EConstr.t -> 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 -> EConstr.t -> EConstr.t -> bool +val dependent_no_evar : Evd.evar_map -> EConstr.t -> EConstr.t -> bool +val dependent_univs : Evd.evar_map -> EConstr.t -> EConstr.t -> bool +val dependent_univs_no_evar : Evd.evar_map -> EConstr.t -> EConstr.t -> bool +val dependent_in_decl : Evd.evar_map -> EConstr.t -> Context.Named.Declaration.t -> bool +val count_occurrences : Evd.evar_map -> EConstr.t -> EConstr.t -> int +val collect_metas : Evd.evar_map -> EConstr.t -> int list +val collect_vars : Evd.evar_map -> EConstr.t -> 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 *) +val occur_term : Evd.evar_map -> EConstr.t -> EConstr.t -> bool (** Synonymous of dependent *) + +(* Substitution of metavariables *) type meta_value_map = (metavariable * constr) list val subst_meta : meta_value_map -> constr -> constr @@ -132,7 +134,7 @@ val subst_meta : meta_value_map -> constr -> constr type meta_type_map = (metavariable * types) list (** [pop c] lifts by -1 the positive indexes in [c] *) -val pop : constr -> constr +val pop : EConstr.t -> constr (** {6 ... } *) (** Substitution of an arbitrary large term. Uses equality modulo @@ -140,20 +142,20 @@ 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 -> EConstr.t -> EConstr.t -> bool) -> EConstr.t -> EConstr.t -> 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) -> - constr -> constr -> constr -> constr + Evd.evar_map -> (Evd.evar_map -> EConstr.t -> EConstr.t -> bool) -> + EConstr.t -> EConstr.t -> EConstr.t -> constr (** [subst_term d c] replaces [d] by [Rel 1] in [c] *) -val subst_term : constr -> constr -> constr +val subst_term : Evd.evar_map -> EConstr.t -> EConstr.t -> 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 -> EConstr.t -> EConstr.t -> EConstr.t -> constr (** Alternative term equalities *) val base_sort_cmp : Reduction.conv_pb -> sorts -> sorts -> bool @@ -165,11 +167,11 @@ val eq_constr : constr -> constr -> bool (* FIXME rename: erases universes*) val eta_reduce_head : constr -> constr (** Flattens application lists *) -val collapse_appl : constr -> constr +val collapse_appl : Evd.evar_map -> EConstr.t -> constr (** Remove recursively the casts around a term i.e. [strip_outer_cast (Cast (Cast ... (Cast c, t) ... ))] is [c]. *) -val strip_outer_cast : constr -> constr +val strip_outer_cast : Evd.evar_map -> EConstr.t -> constr exception CannotFilter @@ -182,24 +184,24 @@ exception CannotFilter type subst = (Context.Rel.t * constr) Evar.Map.t val filtering : Context.Rel.t -> 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 -> EConstr.t -> int * Context.Rel.t * constr +val align_prod_letin : Evd.evar_map -> EConstr.t -> EConstr.t -> Context.Rel.t * 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 -> EConstr.t -> int (** Similar to [nb_lam], but gives the number of products instead *) -val nb_prod : constr -> int +val nb_prod : Evd.evar_map -> EConstr.t -> 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 -> EConstr.t -> int (** Get the last arg of a constr intended to be an application *) -val last_arg : constr -> constr +val last_arg : Evd.evar_map -> EConstr.t -> 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 -> EConstr.t -> constr * constr array val adjust_app_list_size : constr -> constr list -> constr -> constr list -> (constr * constr list * constr * constr list) @@ -250,19 +252,19 @@ val compact_named_context : Context.Named.t -> Context.Compacted.t 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 -> EConstr.t -> Id.t list +val global_vars_set_of_decl : env -> Evd.evar_map -> Context.Named.Declaration.t -> Id.Set.t (** 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 -> Context.Named.t -> 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 isGlobalRef : Evd.evar_map -> EConstr.t -> bool -val is_template_polymorphic : env -> constr -> bool +val is_template_polymorphic : env -> Evd.evar_map -> EConstr.t -> bool (** Combinators on judgments *) diff --git a/ltac/extratactics.ml4 b/ltac/extratactics.ml4 index 8ea60b39a..beaf778a6 100644 --- a/ltac/extratactics.ml4 +++ b/ltac/extratactics.ml4 @@ -747,14 +747,14 @@ let mkCaseEq a : unit Proofview.tactic = let case_eq_intros_rewrite x = Proofview.Goal.nf_enter { enter = begin fun gl -> - let n = nb_prod (Proofview.Goal.concl gl) in + let n = nb_prod (Tacmach.New.project gl) (EConstr.of_constr (Proofview.Goal.concl gl)) in (* Pp.msgnl (Printer.pr_lconstr x); *) Tacticals.New.tclTHENLIST [ mkCaseEq x; Proofview.Goal.nf_enter { enter = begin fun gl -> let concl = Proofview.Goal.concl gl in let hyps = Tacmach.New.pf_ids_of_hyps gl in - let n' = nb_prod concl in + let n' = nb_prod (Tacmach.New.project gl) (EConstr.of_constr concl) in let h = Tacmach.New.of_old (fun g -> fresh_id hyps (Id.of_string "heq") g) gl in Tacticals.New.tclTHENLIST [ Tacticals.New.tclDO (n'-n-1) intro; diff --git a/ltac/rewrite.ml b/ltac/rewrite.ml index 217f5f42e..cd76d4746 100644 --- a/ltac/rewrite.ml +++ b/ltac/rewrite.ml @@ -1527,23 +1527,24 @@ let cl_rewrite_clause_aux ?(abs=None) strat env avoid sigma concl is_hyp : resul in Some (Some (evars, res, newt)) (** Insert a declaration after the last declaration it depends on *) -let rec insert_dependent env decl accu hyps = match hyps with +let rec insert_dependent env sigma decl accu hyps = match hyps with | [] -> List.rev_append accu [decl] | ndecl :: rem -> - if occur_var_in_decl env (NamedDecl.get_id ndecl) decl then + if occur_var_in_decl env sigma (NamedDecl.get_id ndecl) decl then List.rev_append accu (decl :: hyps) else - insert_dependent env decl (ndecl :: accu) rem + insert_dependent env sigma decl (ndecl :: accu) rem let assert_replacing id newt tac = let prf = Proofview.Goal.nf_enter { enter = begin fun gl -> let concl = Proofview.Goal.concl gl in let env = Proofview.Goal.env gl in + let sigma = Tacmach.New.project gl in let ctx = Environ.named_context env in let after, before = List.split_when (NamedDecl.get_id %> Id.equal id) ctx in let nc = match before with | [] -> assert false - | d :: rem -> insert_dependent env (LocalAssum (NamedDecl.get_id d, newt)) [] after @ rem + | d :: rem -> insert_dependent env sigma (LocalAssum (NamedDecl.get_id d, newt)) [] after @ rem in let env' = Environ.reset_with_named_context (val_of_named_context nc) env in Refine.refine ~unsafe:false { run = begin fun sigma -> @@ -1616,7 +1617,7 @@ let cl_rewrite_clause_newtac ?abs ?origsigma ~progress strat clause = | Some id -> (** Only consider variables not depending on [id] *) let ctx = Environ.named_context env in - let filter decl = not (occur_var_in_decl env id decl) in + let filter decl = not (occur_var_in_decl env sigma id decl) in let nctx = List.filter filter ctx in Environ.reset_with_named_context (Environ.val_of_named_context nctx) env in diff --git a/ltac/tauto.ml b/ltac/tauto.ml index 756958c2f..6eab003b5 100644 --- a/ltac/tauto.ml +++ b/ltac/tauto.ml @@ -16,6 +16,7 @@ open Tacexpr open Tacinterp open Util open Tacticals.New +open Proofview.Notations let tauto_plugin = "tauto" let () = Mltop.add_known_module tauto_plugin @@ -111,14 +112,16 @@ let split = Tactics.split_with_bindings false [Misctypes.NoBindings] (** Test *) let is_empty _ ist = - if is_empty_type (assoc_var "X1" ist) then idtac else fail + Proofview.tclEVARMAP >>= fun sigma -> + if is_empty_type sigma (assoc_var "X1" ist) then idtac else fail (* Strictly speaking, this exceeds the propositional fragment as it matches also equality types (and solves them if a reflexivity) *) let is_unit_or_eq _ ist = + Proofview.tclEVARMAP >>= fun sigma -> let flags = assoc_flags ist in let test = if flags.strict_unit then is_unit_type else is_unit_or_eq_type in - if test (assoc_var "X1" ist) then idtac else fail + if test sigma (assoc_var "X1" ist) then idtac else fail let bugged_is_binary t = isApp t && @@ -132,21 +135,23 @@ let bugged_is_binary t = (** Dealing with conjunction *) let is_conj _ ist = + Proofview.tclEVARMAP >>= fun sigma -> let flags = assoc_flags ist in let ind = assoc_var "X1" ist in if (not flags.binary_mode_bugged_detection || bugged_is_binary ind) && - is_conjunction + is_conjunction sigma ~strict:flags.strict_in_hyp_and_ccl ~onlybinary:flags.binary_mode ind then idtac else fail let flatten_contravariant_conj _ ist = + Proofview.tclEVARMAP >>= fun sigma -> let flags = assoc_flags ist in let typ = assoc_var "X1" ist in let c = assoc_var "X2" ist in let hyp = assoc_var "id" ist in - match match_with_conjunction + match match_with_conjunction sigma ~strict:flags.strict_in_contravariant_hyp ~onlybinary:flags.binary_mode typ with @@ -160,21 +165,23 @@ let flatten_contravariant_conj _ ist = (** Dealing with disjunction *) let is_disj _ ist = + Proofview.tclEVARMAP >>= fun sigma -> let flags = assoc_flags ist in let t = assoc_var "X1" ist in if (not flags.binary_mode_bugged_detection || bugged_is_binary t) && - is_disjunction + is_disjunction sigma ~strict:flags.strict_in_hyp_and_ccl ~onlybinary:flags.binary_mode t then idtac else fail let flatten_contravariant_disj _ ist = + Proofview.tclEVARMAP >>= fun sigma -> let flags = assoc_flags ist in let typ = assoc_var "X1" ist in let c = assoc_var "X2" ist in let hyp = assoc_var "id" ist in - match match_with_disjunction + match match_with_disjunction sigma ~strict:flags.strict_in_contravariant_hyp ~onlybinary:flags.binary_mode typ with diff --git a/plugins/btauto/refl_btauto.ml b/plugins/btauto/refl_btauto.ml index 2c5b108e5..3ba5da149 100644 --- a/plugins/btauto/refl_btauto.ml +++ b/plugins/btauto/refl_btauto.ml @@ -14,8 +14,8 @@ let get_inductive dir s = let glob_ref () = Coqlib.find_reference contrib_name ("Coq" :: dir) s in Lazy.from_fun (fun () -> Globnames.destIndRef (glob_ref ())) -let decomp_term (c : Term.constr) = - Term.kind_of_term (Termops.strip_outer_cast c) +let decomp_term sigma (c : Term.constr) = + Term.kind_of_term (Termops.strip_outer_cast sigma (EConstr.of_constr c)) let lapp c v = Term.mkApp (Lazy.force c, v) @@ -105,7 +105,7 @@ module Bool = struct | Negb of t | Ifb of t * t * t - let quote (env : Env.t) (c : Term.constr) : t = + let quote (env : Env.t) sigma (c : Term.constr) : t = let trueb = Lazy.force trueb in let falseb = Lazy.force falseb in let andb = Lazy.force andb in @@ -113,7 +113,7 @@ module Bool = struct let xorb = Lazy.force xorb in let negb = Lazy.force negb in - let rec aux c = match decomp_term c with + let rec aux c = match decomp_term sigma c with | Term.App (head, args) -> if head === andb && Array.length args = 2 then Andb (aux args.(0), aux args.(1)) @@ -181,7 +181,7 @@ module Btauto = struct let var = lapp witness [|p|] in (* Compute an assignment that dissatisfies the goal *) let _, var = Tacmach.pf_reduction_of_red_expr gl (Genredexpr.CbvVm None) var in - let rec to_list l = match decomp_term l with + let rec to_list l = match decomp_term (Tacmach.project gl) l with | Term.App (c, _) when c === (Lazy.force CoqList._nil) -> [] | Term.App (c, [|_; h; t|]) @@ -220,7 +220,7 @@ module Btauto = struct Proofview.Goal.nf_enter { enter = begin fun gl -> let concl = Proofview.Goal.concl gl in let eq = Lazy.force eq in - let t = decomp_term concl in + let t = decomp_term (Tacmach.New.project gl) concl in match t with | Term.App (c, [|typ; p; _|]) when c === eq -> (* should be an equality [@eq poly ?p (Cst false)] *) @@ -234,15 +234,16 @@ module Btauto = struct let tac = Proofview.Goal.nf_enter { enter = begin fun gl -> let concl = Proofview.Goal.concl gl in + let sigma = Tacmach.New.project gl in let eq = Lazy.force eq in let bool = Lazy.force Bool.typ in - let t = decomp_term concl in + let t = decomp_term sigma concl in match t with | Term.App (c, [|typ; tl; tr|]) when typ === bool && c === eq -> let env = Env.empty () in - let fl = Bool.quote env tl in - let fr = Bool.quote env tr in + let fl = Bool.quote env sigma tl in + let fr = Bool.quote env sigma tr in let env = Env.to_list env in let fl = reify env fl in let fr = reify env fr in diff --git a/plugins/cc/cctac.ml b/plugins/cc/cctac.ml index b5ca2f50f..425bb2d6f 100644 --- a/plugins/cc/cctac.ml +++ b/plugins/cc/cctac.ml @@ -58,8 +58,8 @@ let rec decompose_term env sigma t= let tf=decompose_term env sigma f in let targs=Array.map (decompose_term env sigma) args in Array.fold_left (fun s t->Appli (s,t)) tf targs - | Prod (_,a,_b) when not (Termops.dependent (mkRel 1) _b) -> - let b = Termops.pop _b in + | Prod (_,a,_b) when EConstr.Vars.noccurn sigma 1 (EConstr.of_constr _b) -> + let b = Termops.pop (EConstr.of_constr _b) in let sort_b = sf_of env sigma b in let sort_a = sf_of env sigma a in Appli(Appli(Product (sort_a,sort_b) , @@ -86,7 +86,7 @@ let rec decompose_term env sigma t= let p' = Projection.map canon_const p in (Appli (Symb (mkConst (Projection.constant p')), decompose_term env sigma c)) | _ -> - let t = Termops.strip_outer_cast t in + let t = Termops.strip_outer_cast sigma (EConstr.of_constr t) in if closed0 t then Symb t else raise Not_found (* decompose equality in members and type *) @@ -112,8 +112,8 @@ let rec pattern_of_constr env sigma c = (Array.map_to_list (pattern_of_constr env sigma) args) in PApp (pf,List.rev pargs), List.fold_left Int.Set.union Int.Set.empty lrels - | Prod (_,a,_b) when not (Termops.dependent (mkRel 1) _b) -> - let b = Termops.pop _b in + | Prod (_,a,_b) when EConstr.Vars.noccurn sigma 1 (EConstr.of_constr _b) -> + let b = Termops.pop (EConstr.of_constr _b) in let pa,sa = pattern_of_constr env sigma a in let pb,sb = pattern_of_constr env sigma b in let sort_b = sf_of env sigma b in diff --git a/plugins/decl_mode/decl_interp.ml b/plugins/decl_mode/decl_interp.ml index f68c01b18..65d273faf 100644 --- a/plugins/decl_mode/decl_interp.ml +++ b/plugins/decl_mode/decl_interp.ml @@ -374,7 +374,7 @@ let interp_cases info env sigma params (pat:cases_pattern_expr) hyps = match st'.st_it with Thesis nam -> {st_it=Thesis nam;st_label=st'.st_label} | This _ -> {st_it = This st.st_it;st_label=st.st_label} in - let thyps = fst (match_hyps blend nam2 (Termops.pop rest1) hyps) in + let thyps = fst (match_hyps blend nam2 (Termops.pop (EConstr.of_constr rest1)) hyps) in tparams,{pat_vars=tpatvars; pat_aliases=taliases; pat_constr=pat_pat; diff --git a/plugins/decl_mode/decl_proof_instr.ml b/plugins/decl_mode/decl_proof_instr.ml index e19dc86c4..46fa5b408 100644 --- a/plugins/decl_mode/decl_proof_instr.ml +++ b/plugins/decl_mode/decl_proof_instr.ml @@ -445,7 +445,7 @@ let concl_refiner metas body gls = let bsort,_B,nbody = aux nenv (_x::avoid) ((n,mkVar _x)::subst) rest in let body = mkNamedLambda _x _A nbody in - if occur_term (mkVar _x) _B then + if local_occur_var evd _x (EConstr.of_constr _B) then begin let _P = mkNamedLambda _x _A _B in match bsort,sort with @@ -672,14 +672,14 @@ let rec metas_from n hyps = _ :: q -> n :: metas_from (succ n) q | [] -> [] -let rec build_product args body = +let rec build_product sigma args body = match args with (Hprop st| Hvar st )::rest -> - let pprod= lift 1 (build_product rest body) in + let pprod= lift 1 (build_product sigma rest body) in let lbody = match st.st_label with Anonymous -> pprod - | Name id -> subst_term (mkVar id) pprod in + | Name id -> subst_var id pprod in mkProd (st.st_label, st.st_it, lbody) | [] -> body @@ -694,7 +694,7 @@ let instr_suffices _then cut gls0 = let info = get_its_info gls0 in let c_id = pf_get_new_id (Id.of_string "_cofact") gls0 in let ctx,hd = cut.cut_stat in - let c_stat = build_product ctx (mk_stat_or_thesis info gls0 hd) in + let c_stat = build_product (project gls0) ctx (mk_stat_or_thesis info gls0 hd) in let metas = metas_from 1 ctx in let c_ctx,c_head = build_applist c_stat metas in let c_term = applist (mkVar c_id,List.map mkMeta metas) in @@ -780,7 +780,7 @@ let rec consider_match may_intro introduced available expected gls = gls let consider_tac c hyps gls = - match kind_of_term (strip_outer_cast c) with + match kind_of_term (strip_outer_cast (project gls) (EConstr.of_constr c)) with Var id -> consider_match false [] [id] hyps gls | _ -> @@ -805,18 +805,18 @@ let rec take_tac wits gls = (* tactics for define *) -let rec build_function args body = +let rec build_function sigma args body = match args with st::rest -> - let pfun= lift 1 (build_function rest body) in + let pfun= lift 1 (build_function sigma rest body) in let id = match st.st_label with Anonymous -> assert false | Name id -> id in - mkLambda (Name id, st.st_it, subst_term (mkVar id) pfun) + mkLambda (Name id, st.st_it, subst_term sigma (EConstr.mkVar id) (EConstr.of_constr pfun)) | [] -> body let define_tac id args body gls = - let t = build_function args body in + let t = build_function (project gls) args body in Proofview.V82.of_tactic (letin_tac None (Name id) t None Locusops.nowhere) gls (* tactics for reconsider *) @@ -880,7 +880,7 @@ let build_per_info etype casee gls = let concl=pf_concl gls in let env=pf_env gls in let ctyp=pf_unsafe_type_of gls casee in - let is_dep = dependent casee concl in + let is_dep = dependent (project gls) (EConstr.of_constr casee) (EConstr.of_constr concl) in let hd,args = decompose_app (special_whd gls ctyp) in let (ind,u) = try @@ -895,9 +895,9 @@ let build_per_info etype casee gls = let params,real_args = List.chop nparams args in let abstract_obj c body = let typ=pf_unsafe_type_of gls c in - lambda_create env (typ,subst_term c body) in + lambda_create env (typ,subst_term (project gls) (EConstr.of_constr c) (EConstr.of_constr body)) in let pred= List.fold_right abstract_obj - real_args (lambda_create env (ctyp,subst_term casee concl)) in + real_args (lambda_create env (ctyp,subst_term (project gls) (EConstr.of_constr casee) (EConstr.of_constr concl))) in is_dep, {per_casee=casee; per_ctype=ctyp; @@ -953,7 +953,7 @@ let suppose_tac hyps gls0 = let info = get_its_info gls0 in let thesis = pf_concl gls0 in let id = pf_get_new_id (Id.of_string "subcase_") gls0 in - let clause = build_product hyps thesis in + let clause = build_product (project gls0) hyps thesis in let ninfo1 = {pm_stack=Suppose_case::info.pm_stack} in let old_clauses,stack = register_nodep_subcase id info.pm_stack in let ninfo2 = {pm_stack=stack} in @@ -1263,9 +1263,9 @@ let rec execute_cases fix_name per_info tacnext args objs nhrec tree gls = let params,real_args = List.chop nparams all_args in let abstract_obj c body = let typ=pf_unsafe_type_of gls c in - lambda_create env (typ,subst_term c body) in + lambda_create env (typ,subst_term (project gls) (EConstr.of_constr c) (EConstr.of_constr body)) in let elim_pred = List.fold_right abstract_obj - real_args (lambda_create env (ctyp,subst_term casee concl)) in + real_args (lambda_create env (ctyp,subst_term (project gls) (EConstr.of_constr casee) (EConstr.of_constr concl))) in let case_info = Inductiveops.make_case_info env ind RegularStyle in let gen_arities = Inductive.arities_of_constructors (ind,u) spec in let f_ids typ = diff --git a/plugins/extraction/extraction.ml b/plugins/extraction/extraction.ml index a980a43f5..85cacecdb 100644 --- a/plugins/extraction/extraction.ml +++ b/plugins/extraction/extraction.ml @@ -42,11 +42,11 @@ let none = Evd.empty let type_of env c = let polyprop = (lang() == Haskell) in - Retyping.get_type_of ~polyprop env none (strip_outer_cast c) + Retyping.get_type_of ~polyprop env none (strip_outer_cast none (EConstr.of_constr c)) let sort_of env c = let polyprop = (lang() == Haskell) in - Retyping.get_sort_family_of ~polyprop env none (strip_outer_cast c) + Retyping.get_sort_family_of ~polyprop env none (strip_outer_cast none (EConstr.of_constr c)) (*S Generation of flags and signatures. *) @@ -887,7 +887,7 @@ let extract_std_constant env kn body typ = break user's clever let-ins and partial applications). *) let rels, c = let n = List.length s - and m = nb_lam body in + and m = nb_lam Evd.empty (EConstr.of_constr body) (** FIXME *) in if n <= m then decompose_lam_n n body else let s,s' = List.chop m s in diff --git a/plugins/firstorder/formula.ml b/plugins/firstorder/formula.ml index b34a36492..79f185d18 100644 --- a/plugins/firstorder/formula.ml +++ b/plugins/firstorder/formula.ml @@ -79,13 +79,13 @@ type kind_of_formula= let kind_of_formula gl term = let normalize=special_nf gl in let cciterm=special_whd gl term in - match match_with_imp_term cciterm with - Some (a,b)-> Arrow(a,(pop b)) + match match_with_imp_term (project gl) cciterm with + Some (a,b)-> Arrow(a,(pop (EConstr.of_constr b))) |_-> - match match_with_forall_term cciterm with + match match_with_forall_term (project gl) cciterm with Some (_,a,b)-> Forall(a,b) |_-> - match match_with_nodep_ind cciterm with + match match_with_nodep_ind (project gl) cciterm with Some (i,l,n)-> let ind,u=destInd i in let (mib,mip) = Global.lookup_inductive ind in @@ -96,7 +96,7 @@ let kind_of_formula gl term = let has_realargs=(n>0) in let is_trivial= let is_constant c = - Int.equal (nb_prod c) mib.mind_nparams in + Int.equal (nb_prod (project gl) (EConstr.of_constr c)) mib.mind_nparams in Array.exists is_constant mip.mind_nf_lc in if Inductiveops.mis_is_recursive (ind,mib,mip) || (has_realargs && not is_trivial) @@ -108,7 +108,7 @@ let kind_of_formula gl term = else Or((ind,u),l,is_trivial) | _ -> - match match_with_sigma_type cciterm with + match match_with_sigma_type (project gl) cciterm with Some (i,l)-> Exists((destInd i),l) |_-> Atom (normalize cciterm) diff --git a/plugins/firstorder/rules.ml b/plugins/firstorder/rules.ml index 7ffc78928..1d107e9af 100644 --- a/plugins/firstorder/rules.ml +++ b/plugins/firstorder/rules.ml @@ -38,8 +38,8 @@ let wrap n b continue seq gls= []->anomaly (Pp.str "Not the expected number of hyps") | nd::q-> let id = NamedDecl.get_id nd in - if occur_var env id (pf_concl gls) || - List.exists (occur_var_in_decl env id) ctx then + if occur_var env (project gls) id (EConstr.of_constr (pf_concl gls)) || + List.exists (occur_var_in_decl env (project gls) id) ctx then (aux (i-1) q (nd::ctx)) else add_formula Hyp (VarRef id) (NamedDecl.get_type nd) (aux (i-1) q (nd::ctx)) gls in diff --git a/plugins/firstorder/unify.ml b/plugins/firstorder/unify.ml index d9ab36ad6..01c019744 100644 --- a/plugins/firstorder/unify.ml +++ b/plugins/firstorder/unify.ml @@ -21,7 +21,10 @@ exception UFAIL of constr*constr to the equation set. Raises UFAIL with a pair of terms *) +let strip_outer_cast t = strip_outer_cast Evd.empty (EConstr.of_constr t) (** FIXME *) + let unif t1 t2= + let evd = Evd.empty in (** FIXME *) let bige=Queue.create () and sigma=ref [] in let bind i t= @@ -47,18 +50,18 @@ let unif t1 t2= else bind i nt2 | Meta i,_ -> let t=subst_meta !sigma nt2 in - if Int.Set.is_empty (free_rels t) && - not (occur_term (mkMeta i) t) then + if Int.Set.is_empty (free_rels evd (EConstr.of_constr t)) && + not (occur_term evd (EConstr.mkMeta i) (EConstr.of_constr t)) then bind i t else raise (UFAIL(nt1,nt2)) | _,Meta i -> let t=subst_meta !sigma nt1 in - if Int.Set.is_empty (free_rels t) && - not (occur_term (mkMeta i) t) then + if Int.Set.is_empty (free_rels evd (EConstr.of_constr t)) && + not (occur_term evd (EConstr.mkMeta i) (EConstr.of_constr t)) then bind i t else raise (UFAIL(nt1,nt2)) | Cast(_,_,_),_->Queue.add (strip_outer_cast nt1,nt2) bige | _,Cast(_,_,_)->Queue.add (nt1,strip_outer_cast nt2) bige | (Prod(_,a,b),Prod(_,c,d))|(Lambda(_,a,b),Lambda(_,c,d))-> - Queue.add (a,c) bige;Queue.add (pop b,pop d) bige + Queue.add (a,c) bige;Queue.add (pop (EConstr.of_constr b),pop (EConstr.of_constr d)) bige | Case (_,pa,ca,va),Case (_,pb,cb,vb)-> Queue.add (pa,pb) bige; Queue.add (ca,cb) bige; diff --git a/plugins/fourier/fourierR.ml b/plugins/fourier/fourierR.ml index 8e193c753..a14ec8a2c 100644 --- a/plugins/fourier/fourierR.ml +++ b/plugins/fourier/fourierR.ml @@ -461,8 +461,9 @@ exception GoalDone let rec fourier () = Proofview.Goal.nf_enter { enter = begin fun gl -> let concl = Proofview.Goal.concl gl in + let sigma = Tacmach.New.project gl in Coqlib.check_required_library ["Coq";"fourier";"Fourier"]; - let goal = Termops.strip_outer_cast concl in + let goal = Termops.strip_outer_cast sigma (EConstr.of_constr concl) in let fhyp=Id.of_string "new_hyp_for_fourier" in (* si le but est une inéquation, on introduit son contraire, et le but à prouver devient False *) diff --git a/plugins/funind/functional_principles_proofs.ml b/plugins/funind/functional_principles_proofs.ml index 527f4f0b1..f6567ab81 100644 --- a/plugins/funind/functional_principles_proofs.ml +++ b/plugins/funind/functional_principles_proofs.ml @@ -287,7 +287,7 @@ let change_eq env sigma hyp_id (context:Context.Rel.t) x t end_of_type = in let sub = compute_substitution Int.Map.empty (snd t1) (snd t2) in let sub = compute_substitution sub (fst t1) (fst t2) in - let end_of_type_with_pop = Termops.pop end_of_type in (*the equation will be removed *) + let end_of_type_with_pop = Termops.pop (EConstr.of_constr end_of_type) in (*the equation will be removed *) let new_end_of_type = (* Ugly hack to prevent Map.fold order change between ocaml-3.08.3 and ocaml-3.08.4 Can be safely replaced by the next comment for Ocaml >= 3.08.4 @@ -309,7 +309,7 @@ let change_eq env sigma hyp_id (context:Context.Rel.t) x t end_of_type = try let witness = Int.Map.find i sub in if is_local_def decl then anomaly (Pp.str "can not redefine a rel!"); - (Termops.pop end_of_type,ctxt_size,mkLetIn (RelDecl.get_name decl, witness, RelDecl.get_type decl, witness_fun)) + (Termops.pop (EConstr.of_constr end_of_type),ctxt_size,mkLetIn (RelDecl.get_name decl, witness, RelDecl.get_type decl, witness_fun)) with Not_found -> (mkProd_or_LetIn decl end_of_type, ctxt_size + 1, mkLambda_or_LetIn decl witness_fun) ) @@ -430,7 +430,7 @@ let clean_hyp_with_heq ptes_infos eq_hyps hyp_id env sigma = begin let pte,pte_args = (destApp t_x) in let (* fix_info *) prove_rec_hyp = (Id.Map.find (destVar pte) ptes_infos).proving_tac in - let popped_t' = Termops.pop t' in + let popped_t' = Termops.pop (EConstr.of_constr t') in let real_type_of_hyp = it_mkProd_or_LetIn popped_t' context in let prove_new_type_of_hyp = let context_length = List.length context in @@ -480,7 +480,7 @@ let clean_hyp_with_heq ptes_infos eq_hyps hyp_id env sigma = (* observe (str "In "++Ppconstr.pr_id hyp_id++ *) (* str " removing useless precond True" *) (* ); *) - let popped_t' = Termops.pop t' in + let popped_t' = Termops.pop (EConstr.of_constr t') in let real_type_of_hyp = it_mkProd_or_LetIn popped_t' context in @@ -508,7 +508,7 @@ let clean_hyp_with_heq ptes_infos eq_hyps hyp_id env sigma = ] else if is_trivial_eq t_x then (* t_x := t = t => we remove this precond *) - let popped_t' = Termops.pop t' in + let popped_t' = Termops.pop (EConstr.of_constr t') in let real_type_of_hyp = it_mkProd_or_LetIn popped_t' context in @@ -608,7 +608,7 @@ let treat_new_case ptes_infos nb_prod continue_tac term dyn_infos = let fun_body = mkLambda(Anonymous, pf_unsafe_type_of g' term, - Termops.replace_term term (mkRel 1) dyn_infos.info + Termops.replace_term (project g') (EConstr.of_constr term) (EConstr.mkRel 1) (EConstr.of_constr dyn_infos.info) ) in let new_body = pf_nf_betaiota g' (mkApp(fun_body,[| new_term_value |])) in @@ -699,7 +699,7 @@ let build_proof let t = dyn_info'.info in let dyn_infos = {dyn_info' with info = mkCase(ci,ct,t,cb)} in - let g_nb_prod = nb_prod (pf_concl g) in + let g_nb_prod = nb_prod (project g) (EConstr.of_constr (pf_concl g)) in let type_of_term = pf_unsafe_type_of g t in let term_eq = make_refl_eq (Lazy.force refl_equal) type_of_term t @@ -712,7 +712,7 @@ let build_proof (fun g -> observe_tac "toto" ( tclTHENSEQ [Proofview.V82.of_tactic (Simple.case t); (fun g' -> - let g'_nb_prod = nb_prod (pf_concl g') in + let g'_nb_prod = nb_prod (project g') (EConstr.of_constr (pf_concl g')) in let nb_instanciate_partial = g'_nb_prod - g_nb_prod in observe_tac "treat_new_case" (treat_new_case @@ -927,8 +927,8 @@ let generalize_non_dep hyp g = Environ.fold_named_context_reverse (fun (clear,keep) decl -> let hyp = get_id decl in if Id.List.mem hyp hyps - || List.exists (Termops.occur_var_in_decl env hyp) keep - || Termops.occur_var env hyp hyp_typ + || List.exists (Termops.occur_var_in_decl env (project g) hyp) keep + || Termops.occur_var env (project g) hyp (EConstr.of_constr hyp_typ) || Termops.is_section_variable hyp (* should be dangerous *) then (clear,decl::keep) else (hyp::clear,keep)) @@ -1042,7 +1042,7 @@ let do_replace (evd:Evd.evar_map ref) params rec_arg_num rev_args_id f fun_num a let _ = Typing.e_type_of ~refresh:true (Global.env ()) evd res in res in - let nb_intro_to_do = nb_prod (pf_concl g) in + let nb_intro_to_do = nb_prod (project g) (EConstr.of_constr (pf_concl g)) in tclTHEN (tclDO nb_intro_to_do (Proofview.V82.of_tactic intro)) ( diff --git a/plugins/funind/functional_principles_types.ml b/plugins/funind/functional_principles_types.ml index cc699e5d3..032d887de 100644 --- a/plugins/funind/functional_principles_types.ml +++ b/plugins/funind/functional_principles_types.ml @@ -110,7 +110,7 @@ let compute_new_princ_type_from_rel rel_to_fun sorts princ_type = in let dummy_var = mkVar (Id.of_string "________") in let mk_replacement c i args = - let res = mkApp(rel_to_fun.(i), Array.map Termops.pop (array_get_start args)) in + let res = mkApp(rel_to_fun.(i), Array.map (fun c -> Termops.pop (EConstr.of_constr c)) (array_get_start args)) in observe (str "replacing " ++ pr_lconstr c ++ str " by " ++ pr_lconstr res); res in @@ -168,25 +168,25 @@ let compute_new_princ_type_from_rel rel_to_fun sorts princ_type = let new_env = Environ.push_rel (LocalAssum (x,t)) env in let new_b,binders_to_remove_from_b = compute_new_princ_type remove new_env b in if List.exists (eq_constr (mkRel 1)) binders_to_remove_from_b - then (Termops.pop new_b), filter_map (eq_constr (mkRel 1)) Termops.pop binders_to_remove_from_b + then (Termops.pop (EConstr.of_constr new_b)), filter_map (eq_constr (mkRel 1)) (fun c -> Termops.pop (EConstr.of_constr c)) binders_to_remove_from_b else ( bind_fun(new_x,new_t,new_b), list_union_eq eq_constr binders_to_remove_from_t - (List.map Termops.pop binders_to_remove_from_b) + (List.map (fun c -> Termops.pop (EConstr.of_constr c)) binders_to_remove_from_b) ) with | Toberemoved -> (* observe (str "Decl of "++Ppconstr.pr_name x ++ str " is removed "); *) let new_b,binders_to_remove_from_b = compute_new_princ_type remove env (substnl [dummy_var] 1 b) in - new_b, List.map Termops.pop binders_to_remove_from_b + new_b, List.map (fun c -> Termops.pop (EConstr.of_constr c)) binders_to_remove_from_b | Toberemoved_with_rel (n,c) -> (* observe (str "Decl of "++Ppconstr.pr_name x ++ str " is removed "); *) let new_b,binders_to_remove_from_b = compute_new_princ_type remove env (substnl [c] n b) in - new_b, list_add_set_eq eq_constr (mkRel n) (List.map Termops.pop binders_to_remove_from_b) + new_b, list_add_set_eq eq_constr (mkRel n) (List.map (fun c -> Termops.pop (EConstr.of_constr c)) binders_to_remove_from_b) end and compute_new_princ_type_for_letin remove env x v t b = begin @@ -197,25 +197,25 @@ let compute_new_princ_type_from_rel rel_to_fun sorts princ_type = let new_env = Environ.push_rel (LocalDef (x,v,t)) env in let new_b,binders_to_remove_from_b = compute_new_princ_type remove new_env b in if List.exists (eq_constr (mkRel 1)) binders_to_remove_from_b - then (Termops.pop new_b),filter_map (eq_constr (mkRel 1)) Termops.pop binders_to_remove_from_b + then (Termops.pop (EConstr.of_constr new_b)),filter_map (eq_constr (mkRel 1)) (fun c -> Termops.pop (EConstr.of_constr c)) binders_to_remove_from_b else ( mkLetIn(new_x,new_v,new_t,new_b), list_union_eq eq_constr (list_union_eq eq_constr binders_to_remove_from_t binders_to_remove_from_v) - (List.map Termops.pop binders_to_remove_from_b) + (List.map (fun c -> Termops.pop (EConstr.of_constr c)) binders_to_remove_from_b) ) with | Toberemoved -> (* observe (str "Decl of "++Ppconstr.pr_name x ++ str " is removed "); *) let new_b,binders_to_remove_from_b = compute_new_princ_type remove env (substnl [dummy_var] 1 b) in - new_b, List.map Termops.pop binders_to_remove_from_b + new_b, List.map (fun c -> Termops.pop (EConstr.of_constr c)) binders_to_remove_from_b | Toberemoved_with_rel (n,c) -> (* observe (str "Decl of "++Ppconstr.pr_name x ++ str " is removed "); *) let new_b,binders_to_remove_from_b = compute_new_princ_type remove env (substnl [c] n b) in - new_b, list_add_set_eq eq_constr (mkRel n) (List.map Termops.pop binders_to_remove_from_b) + new_b, list_add_set_eq eq_constr (mkRel n) (List.map (fun c -> Termops.pop (EConstr.of_constr c)) binders_to_remove_from_b) end and compute_new_princ_type_with_acc remove env e (c_acc,to_remove_acc) = let new_e,to_remove_from_e = compute_new_princ_type remove env e diff --git a/plugins/funind/indfun.ml b/plugins/funind/indfun.ml index 99b04898b..a264c37c5 100644 --- a/plugins/funind/indfun.ml +++ b/plugins/funind/indfun.ml @@ -14,20 +14,21 @@ open Sigma.Notations module RelDecl = Context.Rel.Declaration -let is_rec_info scheme_info = +let is_rec_info sigma scheme_info = let test_branche min acc decl = acc || ( let new_branche = it_mkProd_or_LetIn mkProp (fst (decompose_prod_assum (RelDecl.get_type decl))) in - let free_rels_in_br = Termops.free_rels new_branche in + let free_rels_in_br = Termops.free_rels sigma (EConstr.of_constr new_branche) in let max = min + scheme_info.Tactics.npredicates in Int.Set.exists (fun i -> i >= min && i< max) free_rels_in_br ) in List.fold_left_i test_branche 1 false (List.rev scheme_info.Tactics.branches) -let choose_dest_or_ind scheme_info = - Tactics.induction_destruct (is_rec_info scheme_info) false +let choose_dest_or_ind scheme_info args = + Proofview.tclBIND Proofview.tclEVARMAP (fun sigma -> + Tactics.induction_destruct (is_rec_info sigma scheme_info) false args) let functional_induction with_clean c princl pat = let res = diff --git a/plugins/funind/invfun.ml b/plugins/funind/invfun.ml index c8b4e4833..cf42a809d 100644 --- a/plugins/funind/invfun.ml +++ b/plugins/funind/invfun.ml @@ -254,7 +254,7 @@ let prove_fun_correct evd functional_induction funs_constr graphs_constr schemes let princ_type = nf_zeta princ_type in let princ_infos = Tactics.compute_elim_sig princ_type in (* The number of args of the function is then easily computable *) - let nb_fun_args = nb_prod (pf_concl g) - 2 in + let nb_fun_args = nb_prod (project g) (EConstr.of_constr (pf_concl g)) - 2 in let args_names = generate_fresh_id (Id.of_string "x") [] nb_fun_args in let ids = args_names@(pf_ids_of_hyps g) in (* Since we cannot ensure that the functional principle is defined in the @@ -467,7 +467,7 @@ let generalize_dependent_of x hyp g = tclMAP (function | LocalAssum (id,t) when not (Id.equal id hyp) && - (Termops.occur_var (pf_env g) x t) -> tclTHEN (Proofview.V82.of_tactic (Tactics.generalize [mkVar id])) (thin [id]) + (Termops.occur_var (pf_env g) (project g) x (EConstr.of_constr t)) -> tclTHEN (Proofview.V82.of_tactic (Tactics.generalize [mkVar id])) (thin [id]) | _ -> tclIDTAC ) (pf_hyps g) @@ -666,7 +666,7 @@ let prove_fun_complete funcs graphs schemes lemmas_types_infos i : tactic = (* Then we get the number of argument of the function and compute a fresh name for each of them *) - let nb_fun_args = nb_prod (pf_concl g) - 2 in + let nb_fun_args = nb_prod (project g) (EConstr.of_constr (pf_concl g)) - 2 in let args_names = generate_fresh_id (Id.of_string "x") [] nb_fun_args in let ids = args_names@(pf_ids_of_hyps g) in (* and fresh names for res H and the principle (cf bug bug #1174) *) @@ -684,7 +684,7 @@ let prove_fun_complete funcs graphs schemes lemmas_types_infos i : tactic = (fun decl -> List.map (fun id -> id) - (generate_fresh_id (Id.of_string "y") ids (nb_prod (RelDecl.get_type decl))) + (generate_fresh_id (Id.of_string "y") ids (nb_prod (project g) (EConstr.of_constr (RelDecl.get_type decl)))) ) branches in diff --git a/plugins/funind/merge.ml b/plugins/funind/merge.ml index 19c2ed417..865042afb 100644 --- a/plugins/funind/merge.ml +++ b/plugins/funind/merge.ml @@ -32,7 +32,7 @@ module RelDecl = Context.Rel.Declaration (** {2 Useful operations on constr and glob_constr} *) -let rec popn i c = if i<=0 then c else pop (popn (i-1) c) +let rec popn i c = if i<=0 then c else pop (EConstr.of_constr (popn (i-1) c)) (** Substitutions in constr *) let compare_constr_nosub t1 t2 = @@ -985,7 +985,7 @@ let relprinctype_to_funprinctype relprinctype nfuns = (* first remove indarg and indarg_in_concl *) let relinfo_noindarg = { relinfo with indarg_in_concl = false; indarg = None; - concl = remove_last_arg (pop relinfo.concl); } in + concl = remove_last_arg (pop (EConstr.of_constr relinfo.concl)); } in (* the nfuns last induction arguments are functional ones: remove them *) let relinfo_argsok = { relinfo_noindarg with nargs = relinfo_noindarg.nargs - nfuns; diff --git a/plugins/funind/recdef.ml b/plugins/funind/recdef.ml index 54066edfb..6b63d7771 100644 --- a/plugins/funind/recdef.ml +++ b/plugins/funind/recdef.ml @@ -407,7 +407,7 @@ let treat_case forbid_new_ids to_intros finalize_tac nb_lam e infos : tactic = let _,args = try destApp ty_teq with DestKO -> assert false in args.(1),args.(2) in - let new_b' = Termops.replace_term teq_lhs teq_rhs new_b in + let new_b' = Termops.replace_term (project g') (EConstr.of_constr teq_lhs) (EConstr.of_constr teq_rhs) (EConstr.of_constr new_b) in let new_infos = { infos with info = new_b'; @@ -681,7 +681,7 @@ let mkDestructEq : (fun decl -> let open Context.Named.Declaration in let id = get_id decl in - if Id.List.mem id not_on_hyp || not (Termops.occur_term expr (get_type decl)) + if Id.List.mem id not_on_hyp || not (Termops.occur_term (project g) (EConstr.of_constr expr) (EConstr.of_constr (get_type decl))) then None else Some id) hyps in let to_revert_constr = List.rev_map mkVar to_revert in let type_of_expr = pf_unsafe_type_of g expr in @@ -1251,7 +1251,7 @@ let clear_goals = | Prod(Name id as na,t',b) -> let b' = clear_goal b in if noccurn 1 b' && (is_rec_res id) - then Termops.pop b' + then Termops.pop (EConstr.of_constr b') else if b' == b then t else mkProd(na,t',b') | _ -> Term.map_constr clear_goal t @@ -1285,7 +1285,7 @@ let open_new_goal build_proof sigma using_lemmas ref_ goal_name (gls_type,decomp anomaly (Pp.str "open_new_goal with an unamed theorem") in let na = next_global_ident_away name [] in - if Termops.occur_existential gls_type then + if Termops.occur_existential sigma (EConstr.of_constr gls_type) then CErrors.error "\"abstract\" cannot handle existentials"; let hook _ _ = let opacity = @@ -1422,7 +1422,7 @@ let start_equation (f:global_reference) (term_f:global_reference) (cont_tactic:Id.t list -> tactic) g = let ids = pf_ids_of_hyps g in let terminate_constr = constr_of_global term_f in - let nargs = nb_prod (fst (type_of_const terminate_constr)) (*FIXME*) in + let nargs = nb_prod (project g) (EConstr.of_constr (fst (type_of_const terminate_constr))) (*FIXME*) in let x = n_x_id ids nargs in observe_tac (str "start_equation") (observe_tclTHENLIST (str "start_equation") [ h_intros x; @@ -1552,7 +1552,7 @@ let recursive_definition is_mes function_name rec_impls type_of_f r rec_arg_num and functional_ref = destConst (constr_of_global functional_ref) and eq_ref = destConst (constr_of_global eq_ref) in generate_induction_principle f_ref tcc_lemma_constr - functional_ref eq_ref rec_arg_num rec_arg_type (nb_prod res) relation; + functional_ref eq_ref rec_arg_num rec_arg_type (nb_prod evm (EConstr.of_constr res)) relation; if Flags.is_verbose () then msgnl (h 1 (Ppconstr.pr_id function_name ++ spc () ++ str"is defined" )++ fnl () ++ diff --git a/plugins/micromega/coq_micromega.ml b/plugins/micromega/coq_micromega.ml index a063cbbfe..49fcf83b4 100644 --- a/plugins/micromega/coq_micromega.ml +++ b/plugins/micromega/coq_micromega.ml @@ -1221,7 +1221,7 @@ struct let g,env,tg = xparse_formula env tg b in mkformula_binary mkIff term f g,env,tg | _ -> parse_atom env tg term) - | Prod(typ,a,b) when not (Termops.dependent (mkRel 1) b)-> + | Prod(typ,a,b) when EConstr.Vars.noccurn Evd.empty 1 (EConstr.of_constr b) (** FIXME *) -> let f,env,tg = xparse_formula env tg a in let g,env,tg = xparse_formula env tg b in mkformula_binary mkI term f g,env,tg @@ -1687,7 +1687,8 @@ let rec mk_topo_order le l = | (Some v,l') -> v :: (mk_topo_order le l') -let topo_sort_constr l = mk_topo_order Termops.dependent l +let topo_sort_constr l = + mk_topo_order (fun c t -> Termops.dependent Evd.empty (** FIXME *) (EConstr.of_constr c) (EConstr.of_constr t)) l (** diff --git a/plugins/quote/quote.ml b/plugins/quote/quote.ml index 6405c8ceb..c6376727a 100644 --- a/plugins/quote/quote.ml +++ b/plugins/quote/quote.ml @@ -183,7 +183,7 @@ type inversion_scheme = { let i_can't_do_that () = error "Quote: not a simple fixpoint" -let decomp_term c = kind_of_term (Termops.strip_outer_cast c) +let decomp_term gl c = kind_of_term (Termops.strip_outer_cast (Tacmach.New.project gl) (EConstr.of_constr c)) (*s [compute_lhs typ i nargsi] builds the term \texttt{(C ?nargsi ... ?2 ?1)}, where \texttt{C} is the [i]-th constructor of inductive @@ -223,14 +223,14 @@ let compute_rhs bodyi index_of_f = let compute_ivs f cs gl = let cst = try destConst f with DestKO -> i_can't_do_that () in let body = Environ.constant_value_in (Global.env()) cst in - match decomp_term body with + match decomp_term gl body with | Fix(([| len |], 0), ([| name |], [| typ |], [| body2 |])) -> let (args3, body3) = decompose_lam body2 in let nargs3 = List.length args3 in let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in let is_conv = Reductionops.is_conv env sigma in - begin match decomp_term body3 with + begin match decomp_term gl body3 with | Case(_,p,c,lci) -> (*

Case c of c1 ... cn end *) let n_lhs_rhs = ref [] and v_lhs = ref (None : constr option) @@ -267,7 +267,7 @@ let compute_ivs f cs gl = (* The Cases predicate is a lambda; we assume no dependency *) let p = match kind_of_term p with - | Lambda (_,_,p) -> Termops.pop p + | Lambda (_,_,p) -> Termops.pop (EConstr.of_constr p) | _ -> p in diff --git a/plugins/rtauto/refl_tauto.ml b/plugins/rtauto/refl_tauto.ml index 367a13333..b129b0bb3 100644 --- a/plugins/rtauto/refl_tauto.ml +++ b/plugins/rtauto/refl_tauto.ml @@ -94,7 +94,7 @@ let rec make_form atom_env gls term = let cciterm=special_whd gls term in match kind_of_term cciterm with Prod(_,a,b) -> - if not (Termops.dependent (mkRel 1) b) && + if EConstr.Vars.noccurn Evd.empty (** FIXME *) 1 (EConstr.of_constr b) && Retyping.get_sort_family_of (pf_env gls) (Tacmach.project gls) a == InProp then @@ -134,7 +134,7 @@ let rec make_hyps atom_env gls lenv = function | LocalAssum (id,typ)::rest -> let hrec= make_hyps atom_env gls (typ::lenv) rest in - if List.exists (Termops.dependent (mkVar id)) lenv || + if List.exists (fun c -> Termops.local_occur_var Evd.empty (** FIXME *) id (EConstr.of_constr c)) lenv || (Retyping.get_sort_family_of (pf_env gls) (Tacmach.project gls) typ != InProp) then diff --git a/plugins/ssrmatching/ssrmatching.ml4 b/plugins/ssrmatching/ssrmatching.ml4 index 77e25b2a5..86cc928c8 100644 --- a/plugins/ssrmatching/ssrmatching.ml4 +++ b/plugins/ssrmatching/ssrmatching.ml4 @@ -702,7 +702,7 @@ let match_upats_HO ~on_instance upats env sigma0 ise c = let fixed_upat = function | {up_k = KpatFlex | KpatEvar _ | KpatProj _} -> false -| {up_t = t} -> not (occur_existential t) +| {up_t = t} -> not (occur_existential Evd.empty (EConstr.of_constr t)) (** FIXME *) let do_once r f = match !r with Some _ -> () | None -> r := Some (f ()) diff --git a/pretyping/cases.ml b/pretyping/cases.ml index d5b125135..6b480986c 100644 --- a/pretyping/cases.ml +++ b/pretyping/cases.ml @@ -535,19 +535,19 @@ let dependencies_in_pure_rhs nargs eqns = let deps_columns = matrix_transpose deps_rows in List.map (List.exists (fun x -> x)) deps_columns -let dependent_decl a = +let dependent_decl sigma a = function - | LocalAssum (na,t) -> dependent a t - | LocalDef (na,c,t) -> dependent a t || dependent a c + | LocalAssum (na,t) -> dependent sigma (EConstr.of_constr a) (EConstr.of_constr t) + | LocalDef (na,c,t) -> dependent sigma (EConstr.of_constr a) (EConstr.of_constr t) || dependent sigma (EConstr.of_constr a) (EConstr.of_constr c) -let rec dep_in_tomatch n = function - | (Pushed _ | Alias _ | NonDepAlias) :: l -> dep_in_tomatch n l - | Abstract (_,d) :: l -> dependent_decl (mkRel n) d || dep_in_tomatch (n+1) l +let rec dep_in_tomatch sigma n = function + | (Pushed _ | Alias _ | NonDepAlias) :: l -> dep_in_tomatch sigma n l + | Abstract (_,d) :: l -> dependent_decl sigma (mkRel n) d || dep_in_tomatch sigma (n+1) l | [] -> false -let dependencies_in_rhs nargs current tms eqns = +let dependencies_in_rhs sigma nargs current tms eqns = match kind_of_term current with - | Rel n when dep_in_tomatch n tms -> List.make nargs true + | Rel n when dep_in_tomatch sigma n tms -> List.make nargs true | _ -> dependencies_in_pure_rhs nargs eqns (* Computing the matrix of dependencies *) @@ -562,24 +562,24 @@ let dependencies_in_rhs nargs current tms eqns = [n-2;1], [1] points to [dn] and [n-2] to [d3] *) -let rec find_dependency_list tmblock = function +let rec find_dependency_list sigma tmblock = function | [] -> [] | (used,tdeps,d)::rest -> - let deps = find_dependency_list tmblock rest in - if used && List.exists (fun x -> dependent_decl x d) tmblock + let deps = find_dependency_list sigma tmblock rest in + if used && List.exists (fun x -> dependent_decl sigma x d) tmblock then List.add_set Int.equal (List.length rest + 1) (List.union Int.equal deps tdeps) else deps -let find_dependencies is_dep_or_cstr_in_rhs (tm,(_,tmtypleaves),d) nextlist = - let deps = find_dependency_list (tm::tmtypleaves) nextlist in +let find_dependencies sigma is_dep_or_cstr_in_rhs (tm,(_,tmtypleaves),d) nextlist = + let deps = find_dependency_list sigma (tm::tmtypleaves) nextlist in if is_dep_or_cstr_in_rhs || not (List.is_empty deps) then ((true ,deps,d)::nextlist) else ((false,[] ,d)::nextlist) -let find_dependencies_signature deps_in_rhs typs = - let l = List.fold_right2 find_dependencies deps_in_rhs typs [] in +let find_dependencies_signature sigma deps_in_rhs typs = + let l = List.fold_right2 (find_dependencies sigma) deps_in_rhs typs [] in List.map (fun (_,deps,_) -> deps) l (* Assume we had terms t1..tq to match in a context xp:Tp,...,x1:T1 |- @@ -1095,30 +1095,30 @@ let rec ungeneralize n ng body = let ungeneralize_branch n k (sign,body) cs = (sign,ungeneralize (n+cs.cs_nargs) k body) -let rec is_dependent_generalization ng body = +let rec is_dependent_generalization sigma ng body = match kind_of_term body with | Lambda (_,_,c) when Int.equal ng 0 -> - dependent (mkRel 1) c + not (EConstr.Vars.noccurn sigma 1 (EConstr.of_constr c)) | Lambda (na,t,c) -> (* We traverse an inner generalization *) - is_dependent_generalization (ng-1) c + is_dependent_generalization sigma (ng-1) c | LetIn (na,b,t,c) -> (* We traverse an alias *) - is_dependent_generalization ng c + is_dependent_generalization sigma ng c | Case (ci,p,c,brs) -> (* We traverse a split *) Array.exists2 (fun q c -> let _,b = decompose_lam_n_decls q c in - is_dependent_generalization ng b) + is_dependent_generalization sigma ng b) ci.ci_cstr_ndecls brs | App (g,args) -> (* We traverse an inner generalization *) assert (isCase g); - is_dependent_generalization (ng+Array.length args) g + is_dependent_generalization sigma (ng+Array.length args) g | _ -> assert false -let is_dependent_branch k (_,br) = - is_dependent_generalization k br +let is_dependent_branch sigma k (_,br) = + is_dependent_generalization sigma k br let postprocess_dependencies evd tocheck brs tomatch pred deps cs = let rec aux k brs tomatch pred tocheck deps = match deps, tomatch with @@ -1126,8 +1126,8 @@ let postprocess_dependencies evd tocheck brs tomatch pred deps cs = | n::deps, Abstract (i,d) :: tomatch -> let d = map_constr (nf_evar evd) d in let is_d = match d with LocalAssum _ -> false | LocalDef _ -> true in - if is_d || List.exists (fun c -> dependent_decl (lift k c) d) tocheck - && Array.exists (is_dependent_branch k) brs then + if is_d || List.exists (fun c -> dependent_decl evd (lift k c) d) tocheck + && Array.exists (is_dependent_branch evd k) brs then (* Dependency in the current term to match and its dependencies is real *) let brs,tomatch,pred,inst = aux (k+1) brs tomatch pred (mkRel n::tocheck) deps in let inst = match d with @@ -1249,8 +1249,8 @@ let build_branch initial current realargs deps (realnames,curname) pb arsign eqn (* We compute over which of x(i+1)..xn and x matching on xi will need a *) (* generalization *) let dep_sign = - find_dependencies_signature - (dependencies_in_rhs const_info.cs_nargs current pb.tomatch eqns) + find_dependencies_signature !(pb.evdref) + (dependencies_in_rhs !(pb.evdref) const_info.cs_nargs current pb.tomatch eqns) (List.rev typs') in (* The dependent term to subst in the types of the remaining UnPushed @@ -1452,7 +1452,7 @@ and compile_alias initial pb (na,orig,(expanded,expanded_typ)) rest = mat = List.map (push_alias_eqn alias) pb.mat } in let j = compile pb in { uj_val = - if isRel c || isVar c || count_occurrences (mkRel 1) j.uj_val <= 1 then + if isRel c || isVar c || count_occurrences !(pb.evdref) (EConstr.mkRel 1) (EConstr.of_constr j.uj_val) <= 1 then subst1 c j.uj_val else mkLetIn (na,c,t,j.uj_val); @@ -1561,7 +1561,7 @@ let matx_of_eqns env eqns = returning True never happens and any inhabited type can be put instead). *) -let adjust_to_extended_env_and_remove_deps env extenv subst t = +let adjust_to_extended_env_and_remove_deps env extenv sigma subst t = let n = Context.Rel.length (rel_context env) in let n' = Context.Rel.length (rel_context extenv) in (* We first remove the bindings that are dependently typed (they are @@ -1583,7 +1583,7 @@ let adjust_to_extended_env_and_remove_deps env extenv subst t = | LocalAssum _ -> p in let p = traverse_local_defs p in let u = lift (n' - n) u in - try Some (p, u, expand_vars_in_term extenv u) + try Some (p, u, expand_vars_in_term extenv sigma u) (* pedrot: does this really happen to raise [Failure _]? *) with Failure _ -> None in let subst0 = List.map_filter map subst in @@ -1617,7 +1617,7 @@ let abstract_tycon loc env evdref subst tycon extenv t = let src = match kind_of_term t with | Evar (evk,_) -> (loc,Evar_kinds.SubEvar evk) | _ -> (loc,Evar_kinds.CasesType true) in - let subst0,t0 = adjust_to_extended_env_and_remove_deps env extenv subst t in + let subst0,t0 = adjust_to_extended_env_and_remove_deps env extenv !evdref subst t in (* We traverse the type T of the original problem Xi looking for subterms that match the non-constructor part of the constraints (this part is in subst); these subterms are the "good" subterms and we replace them @@ -1644,7 +1644,8 @@ let abstract_tycon loc env evdref subst tycon extenv t = let good = List.filter (fun (_,u,_) -> is_conv_leq env !evdref t u) subst in match good with | [] -> - map_constr_with_full_binders push_binder aux x t + let self env c = EConstr.of_constr (aux env (EConstr.Unsafe.to_constr c)) in + EConstr.Unsafe.to_constr (map_constr_with_full_binders !evdref push_binder self x (EConstr.of_constr t)) | (_, _, u) :: _ -> (* u is in extenv *) let vl = List.map pi1 good in let ty = @@ -1652,16 +1653,16 @@ let abstract_tycon loc env evdref subst tycon extenv t = Evarutil.evd_comb1 (refresh_universes (Some false) env) evdref ty in let ty = lift (-k) (aux x ty) in - let depvl = free_rels ty in + let depvl = free_rels !evdref (EConstr.of_constr ty) in let inst = List.map_i (fun i _ -> if Int.List.mem i vl then u else mkRel i) 1 (rel_context extenv) in let rel_filter = - List.map (fun a -> not (isRel a) || dependent a u + List.map (fun a -> not (isRel a) || dependent !evdref (EConstr.of_constr a) (EConstr.of_constr u) || Int.Set.mem (destRel a) depvl) inst in let named_filter = - List.map (fun d -> dependent (mkVar (NamedDecl.get_id d)) u) + List.map (fun d -> local_occur_var !evdref (NamedDecl.get_id d) (EConstr.of_constr u)) (named_context extenv) in let filter = Filter.make (rel_filter @ named_filter) in let candidates = u :: List.map mkRel vl in @@ -1753,7 +1754,7 @@ let build_inversion_problem loc env sigma tms t = List.map (fun (c,d) -> (c,extract_inductive_data pb_env sigma d,d)) decls in let decls = List.rev decls in - let dep_sign = find_dependencies_signature (List.make n true) decls in + let dep_sign = find_dependencies_signature sigma (List.make n true) decls in let sub_tms = List.map2 (fun deps (tm, (tmtyp,_), decl) -> @@ -1878,7 +1879,7 @@ let prepare_predicate_from_arsign_tycon env sigma loc tomatchs arsign c = List.fold_right2 (fun (tm, tmtype) sign (subst, len) -> let signlen = List.length sign in match kind_of_term tm with - | Rel n when dependent tm c + | Rel n when dependent sigma (EConstr.of_constr tm) (EConstr.of_constr c) && Int.equal signlen 1 (* The term to match is not of a dependent type itself *) -> ((n, len) :: subst, len - signlen) | Rel n when signlen > 1 (* The term is of a dependent type, @@ -1890,13 +1891,13 @@ let prepare_predicate_from_arsign_tycon env sigma loc tomatchs arsign c = List.fold_left (fun (subst, len) arg -> match kind_of_term arg with - | Rel n when dependent arg c -> + | Rel n when dependent sigma (EConstr.of_constr arg) (EConstr.of_constr c) -> ((n, len) :: subst, pred len) | _ -> (subst, pred len)) (subst, len) realargs in let subst = - if dependent tm c && List.for_all isRel realargs + if dependent sigma (EConstr.of_constr tm) (EConstr.of_constr c) && List.for_all isRel realargs then (n, len) :: subst else subst in (subst, pred len)) | _ -> (subst, len - signlen)) @@ -2279,7 +2280,7 @@ let lift_ctx n ctx = in ctx' (* Turn matched terms into variables. *) -let abstract_tomatch env tomatchs tycon = +let abstract_tomatch env sigma tomatchs tycon = let prev, ctx, names, tycon = List.fold_left (fun (prev, ctx, names, tycon) (c, t) -> @@ -2288,7 +2289,7 @@ let abstract_tomatch env tomatchs tycon = Rel n -> (lift lenctx c, lift_tomatch_type lenctx t) :: prev, ctx, names, tycon | _ -> let tycon = Option.map - (fun t -> subst_term (lift 1 c) (lift 1 t)) tycon in + (fun t -> subst_term sigma (EConstr.of_constr (lift 1 c)) (EConstr.of_constr (lift 1 t))) tycon in let name = next_ident_away (Id.of_string "filtered_var") names in (mkRel 1, lift_tomatch_type (succ lenctx) t) :: lift_ctx 1 prev, LocalDef (Name name, lift lenctx c, lift lenctx $ type_of_tomatch t) :: ctx, @@ -2406,7 +2407,7 @@ let compile_program_cases loc style (typing_function, evdref) tycon env (* constructors found in patterns *) let tomatchs = coerce_to_indtype typing_function evdref env matx tomatchl in let tycon = valcon_of_tycon tycon in - let tomatchs, tomatchs_lets, tycon' = abstract_tomatch env tomatchs tycon in + let tomatchs, tomatchs_lets, tycon' = abstract_tomatch env !evdref tomatchs tycon in let env = push_rel_context tomatchs_lets env in let len = List.length eqns in let sign, allnames, signlen, eqs, neqs, args = @@ -2460,7 +2461,7 @@ let compile_program_cases loc style (typing_function, evdref) tycon env List.map (fun (c,d) -> (c,extract_inductive_data env !evdref d,d)) typs in let dep_sign = - find_dependencies_signature + find_dependencies_signature !evdref (List.make (List.length typs) true) typs in @@ -2535,7 +2536,7 @@ let compile_cases loc style (typing_fun, evdref) tycon env (predopt, tomatchl, e List.map (fun (c,d) -> (c,extract_inductive_data env sigma d,d)) typs in let dep_sign = - find_dependencies_signature + find_dependencies_signature !evdref (List.make (List.length typs) true) typs in diff --git a/pretyping/coercion.ml b/pretyping/coercion.ml index 2b860ae9c..a3970fc0f 100644 --- a/pretyping/coercion.ml +++ b/pretyping/coercion.ml @@ -487,7 +487,7 @@ let rec inh_conv_coerce_to_fail loc env evd rigidonly v t c1 = let v1 = Option.get v1 in let v2 = Option.map (fun v -> beta_applist (lift 1 v,[v1])) v in let t2 = match v2 with - | None -> subst_term v1 t2 + | None -> subst_term evd' (EConstr.of_constr v1) (EConstr.of_constr t2) | Some v2 -> Retyping.get_type_of env1 evd' v2 in let (evd'',v2') = inh_conv_coerce_to_fail loc env1 evd' rigidonly v2 t2 u2 in (evd'', Option.map (fun v2' -> mkLambda (name, u1, v2')) v2') diff --git a/pretyping/constr_matching.ml b/pretyping/constr_matching.ml index 5ec44a68d..d7b73d333 100644 --- a/pretyping/constr_matching.ml +++ b/pretyping/constr_matching.ml @@ -79,7 +79,7 @@ let add_binders na1 na2 binding_vars (names, terms as subst) = (names, terms) | _ -> subst -let rec build_lambda vars ctx m = match vars with +let rec build_lambda sigma vars ctx m = match vars with | [] -> let len = List.length ctx in lift (-1 * len) m @@ -100,12 +100,12 @@ let rec build_lambda vars ctx m = match vars with let map i = if i > n then pred i else i in let vars = List.map map vars in (** Check that the abstraction is legal *) - let frels = free_rels t in + let frels = free_rels sigma (EConstr.of_constr t) in let brels = List.fold_right Int.Set.add vars Int.Set.empty in let () = if not (Int.Set.subset frels brels) then raise PatternMatchingFailure in (** Create the abstraction *) let m = mkLambda (na, t, m) in - build_lambda vars (pre @ suf) m + build_lambda sigma vars (pre @ suf) m let rec extract_bound_aux k accu frels ctx = match ctx with | [] -> accu @@ -133,12 +133,12 @@ let make_renaming ids = function end | _ -> dummy_constr -let merge_binding allow_bound_rels ctx n cT subst = +let merge_binding sigma allow_bound_rels ctx n cT subst = let c = match ctx with | [] -> (* Optimization *) ([], cT) | _ -> - let frels = free_rels cT in + let frels = free_rels sigma (EConstr.of_constr cT) in if allow_bound_rels then let vars = extract_bound_vars frels ctx in let ordered_vars = Id.Set.elements vars in @@ -169,7 +169,7 @@ let matches_core env sigma convert allow_partial_app allow_bound_rels else false) in let rec sorec ctx env subst p t = - let cT = strip_outer_cast t in + let cT = strip_outer_cast sigma (EConstr.of_constr t) in match p,kind_of_term cT with | PSoApp (n,args),m -> let fold (ans, seen) = function @@ -179,13 +179,13 @@ let matches_core env sigma convert allow_partial_app allow_bound_rels | _ -> error "Only bound indices allowed in second order pattern matching." in let relargs, relset = List.fold_left fold ([], Int.Set.empty) args in - let frels = free_rels cT in + let frels = free_rels sigma (EConstr.of_constr cT) in if Int.Set.subset frels relset then - constrain n ([], build_lambda relargs ctx cT) subst + constrain n ([], build_lambda sigma relargs ctx cT) subst else raise PatternMatchingFailure - | PMeta (Some n), m -> merge_binding allow_bound_rels ctx n cT subst + | PMeta (Some n), m -> merge_binding sigma allow_bound_rels ctx n cT subst | PMeta None, m -> subst @@ -216,7 +216,7 @@ let matches_core env sigma convert allow_partial_app allow_bound_rels let subst = match meta with | None -> subst - | Some n -> merge_binding allow_bound_rels ctx n c subst in + | Some n -> merge_binding sigma allow_bound_rels ctx n c subst in Array.fold_left2 (sorec ctx env) subst args1 args22 else (* Might be a projection on the right *) match kind_of_term c2 with diff --git a/pretyping/detyping.ml b/pretyping/detyping.ml index cad5551c1..72cf31010 100644 --- a/pretyping/detyping.ml +++ b/pretyping/detyping.ml @@ -215,11 +215,11 @@ let lookup_name_as_displayed env t s = | Prod (name,_,c') -> (match compute_displayed_name_in RenamingForGoal avoid name c' with | (Name id,avoid') -> if Id.equal id s then Some n else lookup avoid' (n+1) c' - | (Anonymous,avoid') -> lookup avoid' (n+1) (pop c')) + | (Anonymous,avoid') -> lookup avoid' (n+1) (pop (EConstr.of_constr c'))) | LetIn (name,_,_,c') -> (match compute_displayed_name_in RenamingForGoal avoid name c' with | (Name id,avoid') -> if Id.equal id s then Some n else lookup avoid' (n+1) c' - | (Anonymous,avoid') -> lookup avoid' (n+1) (pop c')) + | (Anonymous,avoid') -> lookup avoid' (n+1) (pop (EConstr.of_constr c'))) | Cast (c,_,_) -> lookup avoid n c | _ -> None in lookup (ids_of_named_context (named_context env)) 1 t @@ -261,13 +261,13 @@ let update_name na ((_,(e,_)),c) = | _ -> na -let rec decomp_branch tags nal b (avoid,env as e) c = +let rec decomp_branch tags nal b (avoid,env as e) sigma c = let flag = if b then RenamingForGoal else RenamingForCasesPattern (fst env,c) in match tags with | [] -> (List.rev nal,(e,c)) | b::tags -> let na,c,f,body,t = - match kind_of_term (strip_outer_cast c), b with + match kind_of_term (strip_outer_cast sigma (EConstr.of_constr c)), b with | Lambda (na,t,c),false -> na,c,compute_displayed_let_name_in,None,Some t | LetIn (na,b,t,c),true -> na,c,compute_displayed_name_in,Some b,Some t @@ -279,17 +279,17 @@ let rec decomp_branch tags nal b (avoid,env as e) c = in let na',avoid' = f flag avoid na c in decomp_branch tags (na'::nal) b - (avoid', add_name_opt na' body t env) c + (avoid', add_name_opt na' body t env) sigma c -let rec build_tree na isgoal e ci cl = +let rec build_tree na isgoal e sigma ci cl = let mkpat n rhs pl = PatCstr(dl,(ci.ci_ind,n+1),pl,update_name na rhs) in let cnl = ci.ci_pp_info.cstr_tags in let cna = ci.ci_cstr_nargs in List.flatten (List.init (Array.length cl) - (fun i -> contract_branch isgoal e (cnl.(i),cna.(i),mkpat i,cl.(i)))) + (fun i -> contract_branch isgoal e sigma (cnl.(i),cna.(i),mkpat i,cl.(i)))) -and align_tree nal isgoal (e,c as rhs) = match nal with +and align_tree nal isgoal (e,c as rhs) sigma = match nal with | [] -> [[],rhs] | na::nal -> match kind_of_term c with @@ -298,20 +298,20 @@ and align_tree nal isgoal (e,c as rhs) = match nal with && not (Int.equal (Array.length cl) 0) && (* don't contract if p dependent *) computable p (List.length ci.ci_pp_info.ind_tags) (* FIXME: can do better *) -> - let clauses = build_tree na isgoal e ci cl in + let clauses = build_tree na isgoal e sigma ci cl in List.flatten (List.map (fun (pat,rhs) -> - let lines = align_tree nal isgoal rhs in + let lines = align_tree nal isgoal rhs sigma in List.map (fun (hd,rest) -> pat::hd,rest) lines) clauses) | _ -> let pat = PatVar(dl,update_name na rhs) in - let mat = align_tree nal isgoal rhs in + let mat = align_tree nal isgoal rhs sigma in List.map (fun (hd,rest) -> pat::hd,rest) mat -and contract_branch isgoal e (cdn,can,mkpat,b) = - let nal,rhs = decomp_branch cdn [] isgoal e b in - let mat = align_tree nal isgoal rhs in +and contract_branch isgoal e sigma (cdn,can,mkpat,b) = + let nal,rhs = decomp_branch cdn [] isgoal e sigma b in + let mat = align_tree nal isgoal rhs sigma in List.map (fun (hd,rhs) -> (mkpat rhs hd,rhs)) mat (**********************************************************************) @@ -439,7 +439,7 @@ let detype_instance sigma l = else Some (List.map (detype_level sigma) (Array.to_list (Univ.Instance.to_array l))) let rec detype flags avoid env sigma t = - match kind_of_term (collapse_appl t) with + match kind_of_term (collapse_appl sigma (EConstr.of_constr t)) with | Rel n -> (try match lookup_name_of_rel n (fst env) with | Name id -> GVar (dl, id) @@ -628,7 +628,7 @@ and share_names flags n l avoid env sigma c t = and detype_eqns flags avoid env sigma ci computable constructs consnargsl bl = try if !Flags.raw_print || not (reverse_matching ()) then raise Exit; - let mat = build_tree Anonymous (snd flags) (avoid,env) ci bl in + let mat = build_tree Anonymous (snd flags) (avoid,env) sigma ci bl in List.map (fun (pat,((avoid,env),c)) -> (dl,[],[pat],detype flags avoid env sigma c)) mat with e when CErrors.noncritical e -> diff --git a/pretyping/evarconv.ml b/pretyping/evarconv.ml index d06009dce..194d0b297 100644 --- a/pretyping/evarconv.ml +++ b/pretyping/evarconv.ml @@ -141,9 +141,10 @@ let check_conv_record env sigma (t1,sk1) (t2,sk2) = match kind_of_term t2 with Prod (_,a,b) -> (* assert (l2=[]); *) let _, a, b = destProd (Evarutil.nf_evar sigma t2) in - if dependent (mkRel 1) b then raise Not_found - else lookup_canonical_conversion (proji, Prod_cs), - (Stack.append_app [|a;pop b|] Stack.empty) + if EConstr.Vars.noccurn sigma 1 (EConstr.of_constr b) then + lookup_canonical_conversion (proji, Prod_cs), + (Stack.append_app [|a;pop (EConstr.of_constr b)|] Stack.empty) + else raise Not_found | Sort s -> lookup_canonical_conversion (proji, Sort_cs (family_of_sort s)),[] @@ -178,7 +179,7 @@ let check_conv_record env sigma (t1,sk1) (t2,sk2) = let c' = subst_univs_level_constr subst c in let t' = subst_univs_level_constr subst t' in let bs' = List.map (subst_univs_level_constr subst) bs in - let h, _ = decompose_app_vect t' in + let h, _ = decompose_app_vect sigma (EConstr.of_constr t') in ctx',(h, t2),c',bs',(Stack.append_app_list params Stack.empty,params1), (Stack.append_app_list us Stack.empty,us2),(extra_args1,extra_args2),c1, (n,Stack.zip(t2,sk2)) @@ -372,7 +373,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty | None -> fallback () | Some l1' -> (* Miller-Pfenning's patterns unification *) let t2 = nf_evar evd tM in - let t2 = solve_pattern_eqn env l1' t2 in + let t2 = solve_pattern_eqn env evd l1' t2 in solve_simple_eqn (evar_conv_x ts) env evd (position_problem on_left pbty,ev,t2) in @@ -893,7 +894,7 @@ and conv_record trs env evd (ctx,(h,h2),c,bs,(params,params1),(us,us2),(sk1,sk2) (fun i -> exact_ise_stack2 env i (evar_conv_x trs) sk1 sk2); test; (fun i -> evar_conv_x trs env i CONV h2 - (fst (decompose_app_vect (substl ks h))))] + (fst (decompose_app_vect i (EConstr.of_constr (substl ks h)))))] else UnifFailure(evd,(*dummy*)NotSameHead) and eta_constructor ts env evd sk1 ((ind, i), u) sk2 term2 = @@ -973,14 +974,14 @@ let apply_on_subterm env evdref f c t = in applyrec (env,(0,c)) t -let filter_possible_projections c ty ctxt args = +let filter_possible_projections evd c ty ctxt args = (* Since args in the types will be replaced by holes, we count the fv of args to have a well-typed filter; don't know how necessary it is however to have a well-typed filter here *) - let fv1 = free_rels (mkApp (c,args)) (* Hack: locally untyped *) in - let fv2 = collect_vars (mkApp (c,args)) in + let fv1 = free_rels evd (EConstr.of_constr (mkApp (c,args))) (* Hack: locally untyped *) in + let fv2 = collect_vars evd (EConstr.of_constr (mkApp (c,args))) in let len = Array.length args in - let tyvars = collect_vars ty in + let tyvars = collect_vars evd (EConstr.of_constr ty) in List.map_i (fun i decl -> let () = assert (i < len) in let a = Array.unsafe_get args i in @@ -1039,7 +1040,7 @@ let second_order_matching ts env_rhs evd (evk,args) argoccs rhs = let t = NamedDecl.get_type decl' in let evs = ref [] in let ty = Retyping.get_type_of env_rhs evd c in - let filter' = filter_possible_projections c ty ctxt args in + let filter' = filter_possible_projections evd c ty ctxt args in (id,t,c,ty,evs,Filter.make filter',occs) :: make_subst (ctxt',l,occsl) | _, _, [] -> [] | _ -> anomaly (Pp.str "Signature or instance are shorter than the occurrences list") in diff --git a/pretyping/evarsolve.ml b/pretyping/evarsolve.ml index bafb009f5..ea3ab17a7 100644 --- a/pretyping/evarsolve.ml +++ b/pretyping/evarsolve.ml @@ -73,7 +73,7 @@ let refresh_universes ?(status=univ_rigid) ?(onlyalg=false) ?(refreshset=false) (** Refresh the types of evars under template polymorphic references *) and refresh_term_evars onevars top t = match kind_of_term (whd_evar !evdref t) with - | App (f, args) when is_template_polymorphic env f -> + | App (f, args) when is_template_polymorphic env !evdref (EConstr.of_constr f) -> let pos = get_polymorphic_positions f in refresh_polymorphic_positions args pos | App (f, args) when top && isEvar f -> @@ -356,14 +356,15 @@ let expansion_of_var aliases x = | [] -> x | a::_ -> a -let rec expand_vars_in_term_using aliases t = match kind_of_term t with +let rec expand_vars_in_term_using sigma aliases t = match kind_of_term t with | Rel _ | Var _ -> normalize_alias aliases t | _ -> - map_constr_with_full_binders - extend_alias expand_vars_in_term_using aliases t + let self aliases c = EConstr.of_constr (expand_vars_in_term_using sigma aliases (EConstr.Unsafe.to_constr c)) in + EConstr.Unsafe.to_constr (map_constr_with_full_binders sigma + extend_alias self aliases (EConstr.of_constr t)) -let expand_vars_in_term env = expand_vars_in_term_using (make_alias_map env) +let expand_vars_in_term env sigma = expand_vars_in_term_using sigma (make_alias_map env) let free_vars_and_rels_up_alias_expansion aliases c = let acc1 = ref Int.Set.empty and acc2 = ref Id.Set.empty in @@ -430,8 +431,8 @@ let constr_list_distinct l = | [] -> true in loop l -let get_actual_deps aliases l t = - if occur_meta_or_existential t then +let get_actual_deps evd aliases l t = + if occur_meta_or_existential evd (EConstr.of_constr t) then (* Probably no restrictions on allowed vars in presence of evars *) l else @@ -460,21 +461,21 @@ let remove_instance_local_defs evd evk args = (* Check if an applied evar "?X[args] l" is a Miller's pattern *) -let find_unification_pattern_args env l t = +let find_unification_pattern_args env evd l t = if List.for_all (fun x -> isRel x || isVar x) l (* common failure case *) then let aliases = make_alias_map env in match (try Some (expand_and_check_vars aliases l) with Exit -> None) with - | Some l as x when constr_list_distinct (get_actual_deps aliases l t) -> x + | Some l as x when constr_list_distinct (get_actual_deps evd aliases l t) -> x | _ -> None else None -let is_unification_pattern_meta env nb m l t = +let is_unification_pattern_meta env evd nb m l t = (* Variables from context and rels > nb are implicitly all there *) (* so we need to be a rel <= nb *) if List.for_all (fun x -> isRel x && destRel x <= nb) l then - match find_unification_pattern_args env l t with - | Some _ as x when not (dependent (mkMeta m) t) -> x + match find_unification_pattern_args env evd l t with + | Some _ as x when not (dependent evd (EConstr.mkMeta m) (EConstr.of_constr t)) -> x | _ -> None else None @@ -485,7 +486,7 @@ let is_unification_pattern_evar env evd (evk,args) l t = then let args = remove_instance_local_defs evd evk args in let n = List.length args in - match find_unification_pattern_args env (args @ l) t with + match find_unification_pattern_args env evd (args @ l) t with | Some l -> Some (List.skipn n l) | _ -> None else None @@ -498,7 +499,7 @@ let is_unification_pattern_pure_evar env evd (evk,args) t = let is_unification_pattern (env,nb) evd f l t = match kind_of_term f with - | Meta m -> is_unification_pattern_meta env nb m l t + | Meta m -> is_unification_pattern_meta env evd nb m l t | Evar ev -> is_unification_pattern_evar env evd ev l t | _ -> None @@ -509,9 +510,9 @@ let is_unification_pattern (env,nb) evd f l t = *implicitly* depend on Vars but lambda abstraction will not reflect this dependency: ?X x = ?1 (?1 is a meta) will return \_.?1 while it should return \y. ?1{x\y} (non constant function if ?1 depends on x) (BB) *) -let solve_pattern_eqn env l c = +let solve_pattern_eqn env sigma l c = let c' = List.fold_right (fun a c -> - let c' = subst_term (lift 1 a) (lift 1 c) in + let c' = subst_term sigma (EConstr.of_constr (lift 1 a)) (EConstr.of_constr (lift 1 c)) in match kind_of_term a with (* Rem: if [a] links to a let-in, do as if it were an assumption *) | Rel n -> @@ -550,7 +551,7 @@ let make_projectable_subst aliases sigma evi args = | LocalAssum (id,c), a::rest -> let a = whd_evar sigma a in let cstrs = - let a',args = decompose_app_vect a in + let a',args = decompose_app_vect sigma (EConstr.of_constr a) in match kind_of_term a' with | Construct cstr -> let l = try Constrmap.find (fst cstr) cstrs with Not_found -> [] in @@ -923,12 +924,12 @@ let invert_invertible_arg fullenv evd aliases k (evk,argsv) args' = let set_of_evctx l = List.fold_left (fun s decl -> Id.Set.add (get_id decl) s) Id.Set.empty l -let filter_effective_candidates evi filter candidates = +let filter_effective_candidates evd evi filter candidates = match filter with | None -> candidates | Some filter -> let ids = set_of_evctx (Filter.filter_list filter (evar_context evi)) in - List.filter (fun a -> Id.Set.subset (collect_vars a) ids) candidates + List.filter (fun a -> Id.Set.subset (collect_vars evd (EConstr.of_constr a)) ids) candidates let filter_candidates evd evk filter candidates_update = let evi = Evd.find_undefined evd evk in @@ -939,7 +940,7 @@ let filter_candidates evd evk filter candidates_update = match candidates with | None -> NoUpdate | Some l -> - let l' = filter_effective_candidates evi filter l in + let l' = filter_effective_candidates evd evi filter l in if List.length l = List.length l' && candidates_update = NoUpdate then NoUpdate else @@ -952,7 +953,7 @@ let closure_of_filter evd evk = function | None -> None | Some filter -> let evi = Evd.find_undefined evd evk in - let vars = collect_vars (Evarutil.nf_evar evd (evar_concl evi)) in + let vars = collect_vars evd (EConstr.of_constr (evar_concl evi)) in let test b decl = b || Idset.mem (get_id decl) vars || match decl with | LocalAssum _ -> @@ -999,7 +1000,7 @@ let do_restrict_hyps evd (evk,args as ev) filter candidates = (* ?e is assumed to have no candidates *) let postpone_non_unique_projection env evd pbty (evk,argsv as ev) sols rhs = - let rhs = expand_vars_in_term env rhs in + let rhs = expand_vars_in_term env evd rhs in let filter = restrict_upon_filter evd evk (* Keep only variables that occur in rhs *) @@ -1010,7 +1011,7 @@ let postpone_non_unique_projection env evd pbty (evk,argsv as ev) sols rhs = (* expands only rels and vars aliases, not rels or vars bound to an *) (* arbitrary complex term *) (fun a -> not (isRel a || isVar a) - || dependent a rhs || List.exists (fun (id,_) -> isVarId id a) sols) + || dependent evd (EConstr.of_constr a) (EConstr.of_constr rhs) || List.exists (fun (id,_) -> isVarId id a) sols) argsv in let filter = closure_of_filter evd evk filter in let candidates = extract_candidates sols in @@ -1060,7 +1061,7 @@ let restrict_candidates conv_algo env evd filter1 (evk1,argsv1) (evk2,argsv2) = | _, None -> filter_candidates evd evk1 filter1 NoUpdate | None, Some _ -> raise DoesNotPreserveCandidateRestriction | Some l1, Some l2 -> - let l1 = filter_effective_candidates evi1 filter1 l1 in + let l1 = filter_effective_candidates evd evi1 filter1 l1 in let l1' = List.filter (fun c1 -> let c1' = instantiate_evar_array evi1 c1 argsv1 in let filter c2 = @@ -1091,9 +1092,7 @@ exception CannotProject of evar_map * existential *) let rec is_constrainable_in top evd k (ev,(fv_rels,fv_ids) as g) t = - let f,args2 = decompose_app_vect t in - let f,args1 = decompose_app_vect (whd_evar evd f) in - let args = Array.append args1 args2 in + let f,args = decompose_app_vect evd (EConstr.of_constr t) in match kind_of_term f with | Construct ((ind,_),u) -> let n = Inductiveops.inductive_nparams ind in @@ -1450,7 +1449,7 @@ let rec invert_definition conv_algo choose env evd pbty (evk,argsv as ev) rhs = | _ -> progress := true; match - let c,args = decompose_app_vect t in + let c,args = decompose_app_vect !evdref (EConstr.of_constr t) in match kind_of_term c with | Construct (cstr,u) when noccur_between 1 k t -> (* This is common case when inferring the return clause of match *) @@ -1466,10 +1465,11 @@ let rec invert_definition conv_algo choose env evd pbty (evk,argsv as ev) rhs = let ty = get_type_of env' !evdref t in let candidates = try + let self env c = EConstr.of_constr (imitate env (EConstr.Unsafe.to_constr c)) in let t = - map_constr_with_full_binders (fun d (env,k) -> push_rel d env, k+1) - imitate envk t in - t::l + map_constr_with_full_binders !evdref (fun d (env,k) -> push_rel d env, k+1) + self envk (EConstr.of_constr t) in + EConstr.Unsafe.to_constr t::l with e when CErrors.noncritical e -> l in (match candidates with | [x] -> x @@ -1480,8 +1480,9 @@ let rec invert_definition conv_algo choose env evd pbty (evk,argsv as ev) rhs = evar'') | None -> (* Evar/Rigid problem (or assimilated if not normal): we "imitate" *) - map_constr_with_full_binders (fun d (env,k) -> push_rel d env, k+1) - imitate envk t + let self env c = EConstr.of_constr (imitate env (EConstr.Unsafe.to_constr c)) in + EConstr.Unsafe.to_constr (map_constr_with_full_binders !evdref (fun d (env,k) -> push_rel d env, k+1) + self envk (EConstr.of_constr t)) in let rhs = whd_beta evd rhs (* heuristic *) in let fast rhs = @@ -1498,7 +1499,7 @@ let rec invert_definition conv_algo choose env evd pbty (evk,argsv as ev) rhs = in is_id_subst filter_ctxt (Array.to_list argsv) && closed0 rhs && - Idset.subset (collect_vars rhs) !names + Idset.subset (collect_vars evd (EConstr.of_constr rhs)) !names in let body = if fast rhs then nf_evar evd rhs @@ -1530,7 +1531,7 @@ and evar_define conv_algo ?(choose=false) env evd pbty (evk,argsv as ev) rhs = with NoCandidates -> try let (evd',body) = invert_definition conv_algo choose env evd pbty ev rhs in - if occur_meta body then raise MetaOccurInBodyInternal; + if occur_meta evd' (EConstr.of_constr body) then raise MetaOccurInBodyInternal; (* invert_definition may have instantiate some evars of rhs with evk *) (* so we recheck acyclicity *) if occur_evar_upto_types evd' evk body then raise (OccurCheckIn (evd',body)); diff --git a/pretyping/evarsolve.mli b/pretyping/evarsolve.mli index f94c83b6d..cf059febf 100644 --- a/pretyping/evarsolve.mli +++ b/pretyping/evarsolve.mli @@ -18,7 +18,7 @@ val is_success : unification_result -> bool (** Replace the vars and rels that are aliases to other vars and rels by their representative that is most ancient in the context *) -val expand_vars_in_term : env -> constr -> constr +val expand_vars_in_term : env -> evar_map -> constr -> constr (** [evar_define choose env ev c] try to instantiate [ev] with [c] (typed in [env]), possibly solving related unification problems, possibly leaving open @@ -62,7 +62,7 @@ val is_unification_pattern_evar : env -> evar_map -> existential -> constr list val is_unification_pattern : env * int -> evar_map -> constr -> constr list -> constr -> constr list option -val solve_pattern_eqn : env -> constr list -> constr -> constr +val solve_pattern_eqn : env -> evar_map -> constr list -> constr -> constr val noccur_evar : env -> evar_map -> Evar.t -> constr -> bool diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml index 8f369a811..9b572f376 100644 --- a/pretyping/pretyping.ml +++ b/pretyping/pretyping.ml @@ -751,7 +751,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre match evar_kind_of_term !evdref resj.uj_val with | App (f,args) -> let f = whd_evar !evdref f in - if is_template_polymorphic env.ExtraEnv.env f then + if is_template_polymorphic env.ExtraEnv.env !evdref (EConstr.of_constr f) then (* Special case for inductive type applications that must be refreshed right away. *) let sigma = !evdref in @@ -1009,7 +1009,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre | VMcast -> let cj = pretype empty_tycon env evdref lvar c in let cty = nf_evar !evdref cj.uj_type and tval = nf_evar !evdref tval in - if not (occur_existential cty || occur_existential tval) then + if not (occur_existential !evdref (EConstr.of_constr cty) || occur_existential !evdref (EConstr.of_constr tval)) then let (evd,b) = Reductionops.vm_infer_conv env.ExtraEnv.env !evdref cty tval in if b then (evdref := evd; cj, tval) else diff --git a/pretyping/recordops.ml b/pretyping/recordops.ml index cda052b79..e897d5f5c 100644 --- a/pretyping/recordops.ml +++ b/pretyping/recordops.ml @@ -171,7 +171,7 @@ let keep_true_projections projs kinds = let filter (p, (_, b)) = if b then Some p else None in List.map_filter filter (List.combine projs kinds) -let cs_pattern_of_constr t = +let cs_pattern_of_constr sigma t = match kind_of_term t with App (f,vargs) -> begin @@ -179,7 +179,7 @@ let cs_pattern_of_constr t = with e when CErrors.noncritical e -> raise Not_found end | Rel n -> Default_cs, Some n, [] - | Prod (_,a,b) when not (Termops.dependent (mkRel 1) b) -> Prod_cs, None, [a; Termops.pop b] + | Prod (_,a,b) when EConstr.Vars.noccurn sigma 1 (EConstr.of_constr b) -> Prod_cs, None, [a; Termops.pop (EConstr.of_constr b)] | Sort s -> Sort_cs (family_of_sort s), None, [] | _ -> begin @@ -217,7 +217,7 @@ let compute_canonical_projections warn (con,ind) = | Some proji_sp -> begin try - let patt, n , args = cs_pattern_of_constr t in + let patt, n , args = cs_pattern_of_constr Evd.empty t (** FIXME *) in ((ConstRef proji_sp, patt, t, n, args) :: l) with Not_found -> let con_pp = Nametab.pr_global_env Id.Set.empty (ConstRef con) diff --git a/pretyping/recordops.mli b/pretyping/recordops.mli index a6a90c751..4a176760c 100644 --- a/pretyping/recordops.mli +++ b/pretyping/recordops.mli @@ -65,7 +65,7 @@ type obj_typ = { o_TCOMPS : constr list } (** ordered *) (** Return the form of the component of a canonical structure *) -val cs_pattern_of_constr : constr -> cs_pattern * int option * constr list +val cs_pattern_of_constr : Evd.evar_map -> constr -> cs_pattern * int option * constr list val pr_cs_pattern : cs_pattern -> Pp.std_ppcmds diff --git a/pretyping/reductionops.ml b/pretyping/reductionops.ml index a85e493ea..820974888 100644 --- a/pretyping/reductionops.ml +++ b/pretyping/reductionops.ml @@ -197,14 +197,14 @@ module Cst_stack = struct (** [best_replace d cst_l c] makes the best replacement for [d] by [cst_l] in [c] *) - let best_replace d cst_l c = + let best_replace sigma d cst_l c = let reconstruct_head = List.fold_left (fun t (i,args) -> mkApp (t,Array.sub args i (Array.length args - i))) in List.fold_right - (fun (cst,params,args) t -> Termops.replace_term - (reconstruct_head d args) - (applist (cst, List.rev params)) - t) cst_l c + (fun (cst,params,args) t -> Termops.replace_term sigma + (EConstr.of_constr (reconstruct_head d args)) + (EConstr.of_constr (applist (cst, List.rev params))) + (EConstr.of_constr t)) cst_l c let pr l = let open Pp in @@ -612,8 +612,9 @@ let safe_meta_value sigma ev = let strong whdfun env sigma t = let rec strongrec env t = - map_constr_with_full_binders push_rel strongrec env (whdfun env sigma t) in - strongrec env t + let t = EConstr.of_constr (whdfun env sigma (EConstr.Unsafe.to_constr t)) in + map_constr_with_full_binders sigma push_rel strongrec env t in + EConstr.Unsafe.to_constr (strongrec env (EConstr.of_constr t)) let local_strong whdfun sigma = let rec strongrec t = Constr.map strongrec (whdfun sigma t) in @@ -712,14 +713,14 @@ let contract_cofix ?env ?reference (bodynum,(names,types,bodies as typedbodies)) substl closure bodies.(bodynum) (** Similar to the "fix" case below *) -let reduce_and_refold_cofix recfun env refold cst_l cofix sk = +let reduce_and_refold_cofix recfun env sigma refold cst_l cofix sk = let raw_answer = let env = if refold then Some env else None in contract_cofix ?env ?reference:(Cst_stack.reference cst_l) cofix in apply_subst (fun x (t,sk') -> let t' = - if refold then Cst_stack.best_replace (mkCoFix cofix) cst_l t else t in + if refold then Cst_stack.best_replace sigma (mkCoFix cofix) cst_l t else t in recfun x (t',sk')) [] refold Cst_stack.empty raw_answer sk @@ -757,7 +758,7 @@ let contract_fix ?env ?reference ((recindices,bodynum),(names,types,bodies as ty replace the fixpoint by the best constant from [cst_l] Other rels are directly substituted by constants "magically found from the context" in contract_fix *) -let reduce_and_refold_fix recfun env refold cst_l fix sk = +let reduce_and_refold_fix recfun env sigma refold cst_l fix sk = let raw_answer = let env = if refold then None else Some env in contract_fix ?env ?reference:(Cst_stack.reference cst_l) fix in @@ -765,7 +766,7 @@ let reduce_and_refold_fix recfun env refold cst_l fix sk = (fun x (t,sk') -> let t' = if refold then - Cst_stack.best_replace (mkFix fix) cst_l t + Cst_stack.best_replace sigma (mkFix fix) cst_l t else t in recfun x (t',sk')) [] refold Cst_stack.empty raw_answer sk @@ -947,7 +948,7 @@ let rec whd_state_gen ?csts ~refold ~tactic_mode flags env sigma = | Rel 1, [] -> let lc = Array.sub cl 0 (napp-1) in let u = if Int.equal napp 1 then f else appvect (f,lc) in - if noccurn 1 u then (pop u,Stack.empty),Cst_stack.empty else fold () + if noccurn 1 u then (pop (EConstr.of_constr u),Stack.empty),Cst_stack.empty else fold () | _ -> fold () else fold () | _ -> fold ()) @@ -974,7 +975,7 @@ let rec whd_state_gen ?csts ~refold ~tactic_mode flags env sigma = |args, (Stack.Fix (f,s',cst_l)::s'') when use_fix -> let x' = Stack.zip(x,args) in let out_sk = s' @ (Stack.append_app [|x'|] s'') in - reduce_and_refold_fix whrec env refold cst_l f out_sk + reduce_and_refold_fix whrec env sigma refold cst_l f out_sk |args, (Stack.Cst (const,curr,remains,s',cst_l) :: s'') -> let x' = Stack.zip(x,args) in begin match remains with @@ -1010,7 +1011,7 @@ let rec whd_state_gen ?csts ~refold ~tactic_mode flags env sigma = if CClosure.RedFlags.red_set flags CClosure.RedFlags.fCOFIX then match Stack.strip_app stack with |args, ((Stack.Case _ |Stack.Proj _)::s') -> - reduce_and_refold_cofix whrec env refold cst_l cofix stack + reduce_and_refold_cofix whrec env sigma refold cst_l cofix stack |_ -> fold () else fold () @@ -1043,7 +1044,7 @@ let local_whd_state_gen flags sigma = | Rel 1, [] -> let lc = Array.sub cl 0 (napp-1) in let u = if Int.equal napp 1 then f else appvect (f,lc) in - if noccurn 1 u then (pop u,Stack.empty) else s + if noccurn 1 u then (pop (EConstr.of_constr u),Stack.empty) else s | _ -> s else s | _ -> s) @@ -1568,10 +1569,10 @@ let meta_reducible_instance evd b = in let metas = Metaset.fold fold fm Metamap.empty in let rec irec u = - let u = whd_betaiota Evd.empty u in + let u = whd_betaiota Evd.empty u (** FIXME *) in match kind_of_term u with - | Case (ci,p,c,bl) when isMeta (strip_outer_cast c) -> - let m = destMeta (strip_outer_cast c) in + | Case (ci,p,c,bl) when EConstr.isMeta evd (EConstr.of_constr (strip_outer_cast evd (EConstr.of_constr c))) -> + let m = destMeta (strip_outer_cast evd (EConstr.of_constr c)) in (match try let g, s = Metamap.find m metas in @@ -1581,8 +1582,8 @@ let meta_reducible_instance evd b = with | Some g -> irec (mkCase (ci,p,g,bl)) | None -> mkCase (ci,irec p,c,Array.map irec bl)) - | App (f,l) when isMeta (strip_outer_cast f) -> - let m = destMeta (strip_outer_cast f) in + | App (f,l) when EConstr.isMeta evd (EConstr.of_constr (strip_outer_cast evd (EConstr.of_constr f))) -> + let m = destMeta (strip_outer_cast evd (EConstr.of_constr f)) in (match try let g, s = Metamap.find m metas in diff --git a/pretyping/reductionops.mli b/pretyping/reductionops.mli index 4cd7a2a86..8dcf5c084 100644 --- a/pretyping/reductionops.mli +++ b/pretyping/reductionops.mli @@ -44,7 +44,7 @@ module Cst_stack : sig val add_args : constr array -> t -> t val add_cst : constr -> t -> t val best_cst : t -> (constr * constr list) option - val best_replace : constr -> t -> constr -> constr + val best_replace : Evd.evar_map -> constr -> t -> constr -> constr val reference : t -> Constant.t option val pr : t -> Pp.std_ppcmds end diff --git a/pretyping/retyping.ml b/pretyping/retyping.ml index 5b67af3e7..ac3b5ef63 100644 --- a/pretyping/retyping.ml +++ b/pretyping/retyping.ml @@ -93,7 +93,7 @@ let retype ?(polyprop=true) sigma = let rec type_of env cstr = match kind_of_term cstr with | Meta n -> - (try strip_outer_cast (Evd.meta_ftype sigma n).Evd.rebus + (try strip_outer_cast sigma (EConstr.of_constr (Evd.meta_ftype sigma n).Evd.rebus) with Not_found -> retype_error (BadMeta n)) | Rel n -> let ty = RelDecl.get_type (lookup_rel n env) in @@ -124,12 +124,13 @@ let retype ?(polyprop=true) sigma = subst1 b (type_of (push_rel (LocalDef (name,b,c1)) env) c2) | Fix ((_,i),(_,tys,_)) -> tys.(i) | CoFix (i,(_,tys,_)) -> tys.(i) - | App(f,args) when is_template_polymorphic env f -> + | App(f,args) when is_template_polymorphic env sigma (EConstr.of_constr f) -> + let f = whd_evar sigma f in let t = type_of_global_reference_knowing_parameters env f args in - strip_outer_cast (subst_type env sigma t (Array.to_list args)) + strip_outer_cast sigma (EConstr.of_constr (subst_type env sigma t (Array.to_list args))) | App(f,args) -> - strip_outer_cast - (subst_type env sigma (type_of env f) (Array.to_list args)) + strip_outer_cast sigma + (EConstr.of_constr (subst_type env sigma (type_of env f) (Array.to_list args))) | Proj (p,c) -> let ty = type_of env c in (try @@ -152,7 +153,8 @@ let retype ?(polyprop=true) sigma = | Prop Pos, (Type u2) -> Type (Univ.sup Univ.type0_univ u2) | Prop Null, (Type _ as s) -> s | Type u1, Type u2 -> Type (Univ.sup u1 u2)) - | App(f,args) when is_template_polymorphic env f -> + | App(f,args) when is_template_polymorphic env sigma (EConstr.of_constr f) -> + let f = whd_evar sigma f in let t = type_of_global_reference_knowing_parameters env f args in sort_of_atomic_type env sigma t args | App(f,args) -> sort_of_atomic_type env sigma (type_of env f) args @@ -168,7 +170,8 @@ let retype ?(polyprop=true) sigma = let s2 = sort_family_of (push_rel (LocalAssum (name,t)) env) c2 in if not (is_impredicative_set env) && s2 == InSet && sort_family_of env t == InType then InType else s2 - | App(f,args) when is_template_polymorphic env f -> + | App(f,args) when is_template_polymorphic env sigma (EConstr.of_constr f) -> + let f = whd_evar sigma f in let t = type_of_global_reference_knowing_parameters env f args in family_of_sort (sort_of_atomic_type env sigma t args) | App(f,args) -> diff --git a/pretyping/tacred.ml b/pretyping/tacred.ml index 7da738508..ff76abe37 100644 --- a/pretyping/tacred.ml +++ b/pretyping/tacred.ml @@ -177,7 +177,7 @@ let eval_table = Summary.ref (Cmap.empty : frozen) ~name:"evaluation" the xp..x1. *) -let check_fix_reversibility labs args ((lv,i),(_,tys,bds)) = +let check_fix_reversibility sigma labs args ((lv,i),(_,tys,bds)) = let n = List.length labs in let nargs = List.length args in if nargs > n then raise Elimconst; @@ -202,7 +202,7 @@ let check_fix_reversibility labs args ((lv,i),(_,tys,bds)) = raise Elimconst; List.iteri (fun i t_i -> if not (Int.List.mem_assoc (i+1) li) then - let fvs = List.map ((+) (i+1)) (Int.Set.elements (free_rels t_i)) in + let fvs = List.map ((+) (i+1)) (Int.Set.elements (free_rels sigma (EConstr.of_constr t_i))) in match List.intersect Int.equal fvs reversible_rels with | [] -> () | _ -> raise Elimconst) @@ -261,7 +261,7 @@ let compute_consteval_direct env sigma ref = let open Context.Rel.Declaration in srec (push_rel (LocalAssum (id,t)) env) (n+1) (t::labs) onlyproj g | Fix fix when not onlyproj -> - (try check_fix_reversibility labs l fix + (try check_fix_reversibility sigma labs l fix with Elimconst -> NotAnElimination) | Case (_,_,d,_) when isRel d && not onlyproj -> EliminationCases n | Case (_,_,d,_) -> srec env n labs true d @@ -1102,13 +1102,13 @@ let fold_one_com com env sigma c = (* Reason first on the beta-iota-zeta normal form of the constant as unfold produces it, so that the "unfold f; fold f" configuration works to refold fix expressions *) - let a = subst_term (clos_norm_flags unfold_side_red env sigma rcom) c in + let a = subst_term sigma (EConstr.of_constr (clos_norm_flags unfold_side_red env sigma rcom)) (EConstr.of_constr c) in if not (eq_constr a c) then subst1 com a else (* Then reason on the non beta-iota-zeta form for compatibility - even if it is probably a useless configuration *) - let a = subst_term rcom c in + let a = subst_term sigma (EConstr.of_constr rcom) (EConstr.of_constr c) in subst1 com a let fold_commands cl env sigma c = @@ -1133,8 +1133,8 @@ let compute = cbv_betadeltaiota let abstract_scheme env (locc,a) (c, sigma) = let ta = Retyping.get_type_of env sigma a in let na = named_hd env ta Anonymous in - if occur_meta ta then error "Cannot find a type for the generalisation."; - if occur_meta a then + if occur_meta sigma (EConstr.of_constr ta) then error "Cannot find a type for the generalisation."; + if occur_meta sigma (EConstr.of_constr a) then mkLambda (na,ta,c), sigma else let c', sigma' = subst_closed_term_occ env sigma (AtOccs locc) a c in diff --git a/pretyping/unification.ml b/pretyping/unification.ml index e2b3af7e9..3134dac6a 100644 --- a/pretyping/unification.ml +++ b/pretyping/unification.ml @@ -87,7 +87,7 @@ let abstract_scheme env evd c l lname_typ = are unclear... if occur_meta ta then error "cannot find a type for the generalisation" else *) - if occur_meta a then mkLambda_name env (na,ta,t), evd + if occur_meta evd (EConstr.of_constr a) then mkLambda_name env (na,ta,t), evd else let t', evd' = Find_subterm.subst_closed_term_occ env evd locc a t in mkLambda_name env (na,ta,t'), evd') @@ -182,7 +182,7 @@ let solve_pattern_eqn_array (env,nb) f l c (sigma,metasubst,evarsubst) = extra assumptions added by unification to the context *) let env' = pop_rel_context nb env in let sigma,c = pose_all_metas_as_evars env' sigma c in - let c = solve_pattern_eqn env l c in + let c = solve_pattern_eqn env sigma l c in let pb = (Conv,TypeNotProcessed) in if noccur_between 1 nb c then sigma,(k,lift (-nb) c,pb)::metasubst,evarsubst @@ -190,7 +190,7 @@ let solve_pattern_eqn_array (env,nb) f l c (sigma,metasubst,evarsubst) = | Evar ev -> let env' = pop_rel_context nb env in let sigma,c = pose_all_metas_as_evars env' sigma c in - sigma,metasubst,(env,ev,solve_pattern_eqn env l c)::evarsubst + sigma,metasubst,(env,ev,solve_pattern_eqn env sigma l c)::evarsubst | _ -> assert false let push d (env,n) = (push_rel_assum d env,n+1) @@ -679,7 +679,7 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb if k2 < k1 then sigma,(k1,cN,stN)::metasubst,evarsubst else sigma,(k2,cM,stM)::metasubst,evarsubst | Meta k, _ - when not (dependent cM cN) (* helps early trying alternatives *) -> + when not (dependent sigma (EConstr.of_constr cM) (EConstr.of_constr cN)) (* helps early trying alternatives *) -> let sigma = if opt.with_types && flags.check_applied_meta_types then (try @@ -699,7 +699,7 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb evarsubst) else error_cannot_unify_local curenv sigma (m,n,cN) | _, Meta k - when not (dependent cN cM) (* helps early trying alternatives *) -> + when not (dependent sigma (EConstr.of_constr cN) (EConstr.of_constr cM)) (* helps early trying alternatives *) -> let sigma = if opt.with_types && flags.check_applied_meta_types then (try @@ -728,15 +728,15 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb sigma,metasubst,((curenv,ev,cN)::evarsubst) | Evar (evk,_ as ev), _ when not (Evar.Set.mem evk flags.frozen_evars) - && not (occur_evar evk cN) -> - let cmvars = free_rels cM and cnvars = free_rels cN in + && not (occur_evar sigma evk (EConstr.of_constr cN)) -> + let cmvars = free_rels sigma (EConstr.of_constr cM) and cnvars = free_rels sigma (EConstr.of_constr cN) in if Int.Set.subset cnvars cmvars then sigma,metasubst,((curenv,ev,cN)::evarsubst) else error_cannot_unify_local curenv sigma (m,n,cN) | _, Evar (evk,_ as ev) when not (Evar.Set.mem evk flags.frozen_evars) - && not (occur_evar evk cM) -> - let cmvars = free_rels cM and cnvars = free_rels cN in + && not (occur_evar sigma evk (EConstr.of_constr cM)) -> + let cmvars = free_rels sigma (EConstr.of_constr cM) and cnvars = free_rels sigma (EConstr.of_constr cN) in if Int.Set.subset cmvars cnvars then sigma,metasubst,((curenv,ev,cM)::evarsubst) else error_cannot_unify_local curenv sigma (m,n,cN) @@ -1295,8 +1295,8 @@ let w_merge env with_types flags (evd,metas,evars) = (* This can make rhs' ill-typed if metas are *) let rhs' = subst_meta_instances metas rhs in match kind_of_term rhs with - | App (f,cl) when occur_meta rhs' -> - if occur_evar evk rhs' then + | App (f,cl) when occur_meta evd (EConstr.of_constr rhs') -> + if occur_evar evd evk (EConstr.of_constr rhs') then error_occur_check curenv evd evk rhs'; if is_mimick_head flags.modulo_delta f then let evd' = @@ -1474,16 +1474,16 @@ let iter_fail f a = (* make_abstraction: a variant of w_unify_to_subterm which works on contexts, with evars, and possibly with occurrences *) -let indirectly_dependent c d decls = +let indirectly_dependent sigma c d decls = not (isVar c) && (* This test is not needed if the original term is a variable, but it is needed otherwise, as e.g. when abstracting over "2" in "forall H:0=2, H=H:>(0=1+1) -> 0=2." where there is now obvious way to see that the second hypothesis depends indirectly over 2 *) - List.exists (fun d' -> dependent_in_decl (mkVar (NamedDecl.get_id d')) d) decls + List.exists (fun d' -> dependent_in_decl sigma (EConstr.mkVar (NamedDecl.get_id d')) d) decls -let indirect_dependency d decls = - decls |> List.filter (fun d' -> dependent_in_decl (mkVar (NamedDecl.get_id d')) d) |> List.hd |> NamedDecl.get_id +let indirect_dependency sigma d decls = + decls |> List.filter (fun d' -> dependent_in_decl sigma (EConstr.mkVar (NamedDecl.get_id d')) d) |> List.hd |> NamedDecl.get_id let finish_evar_resolution ?(flags=Pretyping.all_and_fail_flags) env current_sigma (pending,c) = let current_sigma = Sigma.to_evar_map current_sigma in @@ -1610,7 +1610,7 @@ let make_abstraction_core name (test,out) env sigma c ty occs check_occs concl = let occ = if likefirst then LikeFirst else AtOccs occ in let newdecl = replace_term_occ_decl_modulo occ test mkvarid d in if Context.Named.Declaration.equal d newdecl - && not (indirectly_dependent c d depdecls) + && not (indirectly_dependent sigma c d depdecls) then if check_occs && not (in_every_hyp occs) then raise (PretypeError (env,sigma,NoOccurrenceFound (c,Some hyp))) @@ -1695,13 +1695,13 @@ let w_unify_to_subterm env evd ?(flags=default_unify_flags ()) (op,cl) = let bestexn = ref None in let kop = Keys.constr_key op in let rec matchrec cl = - let cl = strip_outer_cast cl in + let cl = strip_outer_cast evd (EConstr.of_constr cl) in (try if closed0 cl && not (isEvar cl) && keyed_unify env evd kop cl then (try if !keyed_unification then - let f1, l1 = decompose_app_vect op in - let f2, l2 = decompose_app_vect cl in + let f1, l1 = decompose_app_vect evd (EConstr.of_constr op) in + let f2, l2 = decompose_app_vect evd (EConstr.of_constr cl) in w_typed_unify_array env evd flags f1 l1 f2 l2,cl else w_typed_unify env evd CONV flags op cl,cl with ex when Pretype_errors.unsatisfiable_exception ex -> @@ -1788,7 +1788,7 @@ let w_unify_to_subterm_all env evd ?(flags=default_unify_flags ()) (op,cl) = in ffail 0 in let rec matchrec cl = - let cl = strip_outer_cast cl in + let cl = strip_outer_cast evd (EConstr.of_constr cl) in (bind (if closed0 cl then return (fun () -> w_typed_unify env evd CONV flags op cl,cl) @@ -1839,7 +1839,7 @@ let w_unify_to_subterm_list env evd flags hdmeta oplist t = else let allow_K = flags.allow_K_in_toplevel_higher_order_unification in let flags = - if occur_meta_or_existential op || !keyed_unification then + if occur_meta_or_existential evd (EConstr.of_constr op) || !keyed_unification then (* This is up to delta for subterms w/o metas ... *) flags else @@ -1848,7 +1848,7 @@ let w_unify_to_subterm_list env evd flags hdmeta oplist t = unify pre-existing non frozen evars of the goal or of the pattern *) set_no_delta_flags flags in - let t' = (strip_outer_cast op,t) in + let t' = (strip_outer_cast evd (EConstr.of_constr op),t) in let (evd',cl) = try if is_keyed_unification () then @@ -1864,7 +1864,7 @@ let w_unify_to_subterm_list env evd flags hdmeta oplist t = (* w_unify_to_subterm does not go through evars, so the next step, which was already in <= 8.4, is needed at least for compatibility of rewrite *) - dependent op t -> (evd,op) + dependent evd (EConstr.of_constr op) (EConstr.of_constr t) -> (evd,op) in if not allow_K && (* ensure we found a different instance *) diff --git a/printing/printer.ml b/printing/printer.ml index 3c8ad4255..b36df27ed 100644 --- a/printing/printer.ml +++ b/printing/printer.ml @@ -760,7 +760,8 @@ let pr_prim_rule = function str ";[" ++ cl ++ str"intro " ++ pr_id id ++ str"|idtac]") | Refine c -> - str(if Termops.occur_meta c then "refine " else "exact ") ++ + (** FIXME *) + str(if Termops.occur_meta Evd.empty (EConstr.of_constr c) then "refine " else "exact ") ++ Constrextern.with_meta_as_hole pr_constr c (* Backwards compatibility *) diff --git a/proofs/clenv.ml b/proofs/clenv.ml index fad656223..34bc83097 100644 --- a/proofs/clenv.ml +++ b/proofs/clenv.ml @@ -77,7 +77,7 @@ let clenv_push_prod cl = | Cast (t,_,_) -> clrec t | Prod (na,t,u) -> let mv = new_meta () in - let dep = dependent (mkRel 1) u in + let dep = not (EConstr.Vars.noccurn (cl_sigma cl) 1 (EConstr.of_constr u)) in let na' = if dep then na else Anonymous in let e' = meta_declare mv t ~name:na' cl.evd in let concl = if dep then subst1 (mkMeta mv) u else u in @@ -332,7 +332,7 @@ let clenv_pose_metas_as_evars clenv dep_mvs = let ty = clenv_meta_type clenv mv in (* Postpone the evar-ization if dependent on another meta *) (* This assumes no cycle in the dependencies - is it correct ? *) - if occur_meta ty then fold clenv (mvs@[mv]) + if occur_meta clenv.evd (EConstr.of_constr ty) then fold clenv (mvs@[mv]) else let src = evar_source_of_meta mv clenv.evd in let src = adjust_meta_source clenv.evd mv src in @@ -404,7 +404,7 @@ type arg_bindings = constr explicit_bindings * of cval, ctyp. *) let clenv_independent clenv = - let mvs = collect_metas (clenv_value clenv) in + let mvs = collect_metas clenv.evd (EConstr.of_constr (clenv_value clenv)) in let ctyp_mvs = (mk_freelisted (clenv_type clenv)).freemetas in let deps = Metaset.union (dependent_in_type_of_metas clenv mvs) ctyp_mvs in List.filter (fun mv -> not (Metaset.mem mv deps)) mvs @@ -522,7 +522,7 @@ let clenv_match_args bl clenv = exception NoSuchBinding let clenv_constrain_last_binding c clenv = - let all_mvs = collect_metas clenv.templval.rebus in + let all_mvs = collect_metas clenv.evd (EConstr.of_constr clenv.templval.rebus) in let k = try List.last all_mvs with Failure _ -> raise NoSuchBinding in clenv_assign_binding clenv k c @@ -612,7 +612,7 @@ let make_evar_clause env sigma ?len t = let sigma = Sigma.Unsafe.of_evar_map sigma in let Sigma (ev, sigma, _) = new_evar ~store env sigma t1 in let sigma = Sigma.to_evar_map sigma in - let dep = dependent (mkRel 1) t2 in + let dep = not (EConstr.Vars.noccurn sigma 1 (EConstr.of_constr t2)) in let hole = { hole_evar = ev; hole_type = t1; @@ -682,9 +682,9 @@ let solve_evar_clause env sigma hyp_only clause = function if h.hole_deps then (** Some subsequent term uses the hole *) let (ev, _) = destEvar h.hole_evar in - let is_dep hole = occur_evar ev hole.hole_type in + let is_dep hole = occur_evar sigma ev (EConstr.of_constr hole.hole_type) in let in_hyp = List.exists is_dep holes in - let in_ccl = occur_evar ev clause.cl_concl in + let in_ccl = occur_evar sigma ev (EConstr.of_constr clause.cl_concl) in let dep = if hyp_only then in_hyp && not in_ccl else in_hyp || in_ccl in let h = { h with hole_deps = dep } in h :: holes diff --git a/proofs/clenvtac.ml b/proofs/clenvtac.ml index 98b5bc8b0..d8a20e08b 100644 --- a/proofs/clenvtac.ml +++ b/proofs/clenvtac.ml @@ -33,12 +33,12 @@ let clenv_cast_meta clenv = | _ -> map_constr crec u and crec_hd u = - match kind_of_term (strip_outer_cast u) with + match kind_of_term (strip_outer_cast clenv.evd (EConstr.of_constr u)) with | Meta mv -> (try let b = Typing.meta_type clenv.evd mv in - assert (not (occur_meta b)); - if occur_meta b then u + assert (not (occur_meta clenv.evd (EConstr.of_constr b))); + if occur_meta clenv.evd (EConstr.of_constr b) then u else mkCast (mkMeta mv, DEFAULTcast, b) with Not_found -> u) | App(f,args) -> mkApp (crec_hd f, Array.map crec args) diff --git a/proofs/evar_refiner.ml b/proofs/evar_refiner.ml index ff0df9179..d4b308bbe 100644 --- a/proofs/evar_refiner.ml +++ b/proofs/evar_refiner.ml @@ -25,7 +25,7 @@ let depends_on_evar evk _ (pbty,_,t1,t2) = with NoHeadEvar -> false let define_and_solve_constraints evk c env evd = - if Termops.occur_evar evk c then + if Termops.occur_evar evd evk (EConstr.of_constr c) then Pretype_errors.error_occur_check env evd evk c; let evd = define evk c evd in let (evd,pbs) = extract_changed_conv_pbs evd (depends_on_evar evk) in diff --git a/proofs/logic.ml b/proofs/logic.ml index 44c629484..29f295682 100644 --- a/proofs/logic.ml +++ b/proofs/logic.ml @@ -138,12 +138,12 @@ let find_q x (m,q) = else find (Id.Set.union l accs) (i::acc) itl in find Id.Set.empty [] q -let occur_vars_in_decl env hyps d = +let occur_vars_in_decl env sigma hyps d = if Id.Set.is_empty hyps then false else - let ohyps = global_vars_set_of_decl env d in + let ohyps = global_vars_set_of_decl env sigma d in Id.Set.exists (fun h -> Id.Set.mem h ohyps) hyps -let reorder_context env sign ord = +let reorder_context env sigma sign ord = let ords = List.fold_right Id.Set.add ord Id.Set.empty in if not (Int.equal (List.length ord) (Id.Set.cardinal ords)) then error "Order list has duplicates"; @@ -152,13 +152,13 @@ let reorder_context env sign ord = | [] -> List.rev ctxt_tail @ ctxt_head | top::ord' when mem_q top moved_hyps -> let ((d,h),mh) = find_q top moved_hyps in - if occur_vars_in_decl env h d then + if occur_vars_in_decl env sigma h d then user_err ~hdr:"reorder_context" (str "Cannot move declaration " ++ pr_id top ++ spc() ++ str "before " ++ pr_sequence pr_id (Id.Set.elements (Id.Set.inter h - (global_vars_set_of_decl env d)))); + (global_vars_set_of_decl env sigma d)))); step ord' expected ctxt_head mh (d::ctxt_tail) | _ -> (match ctxt_head with @@ -173,16 +173,16 @@ let reorder_context env sign ord = ctxt (push_val x moved_hyps) (d::ctxt_tail)) in step ord ords sign mt_q [] -let reorder_val_context env sign ord = - val_of_named_context (reorder_context env (named_context_of_val sign) ord) +let reorder_val_context env sigma sign ord = + val_of_named_context (reorder_context env sigma (named_context_of_val sign) ord) -let check_decl_position env sign d = +let check_decl_position env sigma sign d = let x = NamedDecl.get_id d in - let needed = global_vars_set_of_decl env d in - let deps = dependency_closure env (named_context_of_val sign) needed in + let needed = global_vars_set_of_decl env sigma d in + let deps = dependency_closure env sigma (named_context_of_val sign) needed in if Id.List.mem x deps then user_err ~hdr:"Logic.check_decl_position" (str "Cannot create self-referring hypothesis " ++ pr_id x); @@ -233,12 +233,12 @@ let hyp_of_move_location = function | MoveBefore id -> id | _ -> assert false -let move_hyp toleft (left,declfrom,right) hto = +let move_hyp sigma toleft (left,declfrom,right) hto = let env = Global.env() in let test_dep d d2 = if toleft - then occur_var_in_decl env (NamedDecl.get_id d2) d - else occur_var_in_decl env (NamedDecl.get_id d) d2 + then occur_var_in_decl env sigma (NamedDecl.get_id d2) d + else occur_var_in_decl env sigma (NamedDecl.get_id d) d2 in let rec moverec first middle = function | [] -> @@ -278,10 +278,10 @@ let move_hyp toleft (left,declfrom,right) hto = List.fold_left (fun sign d -> push_named_context_val d sign) right left -let move_hyp_in_named_context hfrom hto sign = +let move_hyp_in_named_context sigma hfrom hto sign = let (left,right,declfrom,toleft) = split_sign hfrom hto (named_context_of_val sign) in - move_hyp toleft (left,declfrom,right) hto + move_hyp sigma toleft (left,declfrom,right) hto (**********************************************************************) @@ -320,10 +320,10 @@ let check_conv_leq_goal env sigma arg ty conclty = else sigma exception Stop of constr list -let meta_free_prefix a = +let meta_free_prefix sigma a = try let _ = Array.fold_left (fun acc a -> - if occur_meta a then raise (Stop acc) + if occur_meta sigma (EConstr.of_constr a) then raise (Stop acc) else a :: acc) [] a in a with Stop acc -> Array.rev_of_list acc @@ -338,7 +338,7 @@ let rec mk_refgoals sigma goal goalacc conclty trm = let mk_goal hyps concl = Goal.V82.mk_goal sigma hyps concl (Goal.V82.extra sigma goal) in - if (not !check) && not (occur_meta trm) then + if (not !check) && not (occur_meta sigma (EConstr.of_constr trm)) then let t'ty = Retyping.get_type_of env sigma trm in let sigma = check_conv_leq_goal env sigma trm t'ty conclty in (goalacc,t'ty,sigma,trm) @@ -346,7 +346,7 @@ let rec mk_refgoals sigma goal goalacc conclty trm = match kind_of_term trm with | Meta _ -> let conclty = nf_betaiota sigma conclty in - if !check && occur_meta conclty then + if !check && occur_meta sigma (EConstr.of_constr conclty) then raise (RefinerError (MetaInType conclty)); let (gl,ev,sigma) = mk_goal hyps conclty in gl::goalacc, conclty, sigma, ev @@ -367,10 +367,10 @@ let rec mk_refgoals sigma goal goalacc conclty trm = | App (f,l) -> let (acc',hdty,sigma,applicand) = - if is_template_polymorphic env f then + if is_template_polymorphic env sigma (EConstr.of_constr f) then let ty = (* Template sort-polymorphism of definition and inductive types *) - let firstmeta = Array.findi (fun i x -> occur_meta x) l in + let firstmeta = Array.findi (fun i x -> occur_meta sigma (EConstr.of_constr x)) l in let args, _ = Option.cata (fun i -> CArray.chop i l) (l, [||]) firstmeta in type_of_global_reference_knowing_parameters env sigma f args in @@ -406,7 +406,7 @@ let rec mk_refgoals sigma goal goalacc conclty trm = (acc'',conclty',sigma, ans) | _ -> - if occur_meta trm then + if occur_meta sigma (EConstr.of_constr trm) then anomaly (Pp.str "refiner called with a meta in non app/case subterm"); let t'ty = goal_type_of env sigma trm in let sigma = check_conv_leq_goal env sigma trm t'ty conclty in @@ -432,9 +432,9 @@ and mk_hdgoals sigma goal goalacc trm = | App (f,l) -> let (acc',hdty,sigma,applicand) = - if is_template_polymorphic env f + if is_template_polymorphic env sigma (EConstr.of_constr f) then - let l' = meta_free_prefix l in + let l' = meta_free_prefix sigma l in (goalacc,type_of_global_reference_knowing_parameters env sigma f l',sigma,f) else mk_hdgoals sigma goal goalacc f in @@ -464,7 +464,7 @@ and mk_hdgoals sigma goal goalacc trm = (acc',ty,sigma,c) | _ -> - if !check && occur_meta trm then + if !check && occur_meta sigma (EConstr.of_constr trm) then anomaly (Pp.str "refine called with a dependent meta"); goalacc, goal_type_of env sigma trm, sigma, trm @@ -511,9 +511,9 @@ let convert_hyp check sign sigma d = if check && not (Option.equal (is_conv env sigma) b c) then user_err ~hdr:"Logic.convert_hyp" (str "Incorrect change of the body of "++ pr_id id ++ str "."); - if check then reorder := check_decl_position env sign d; + if check then reorder := check_decl_position env sigma sign d; d) in - reorder_val_context env sign' !reorder + reorder_val_context env sigma sign' !reorder @@ -537,7 +537,7 @@ let prim_refiner r sigma goal = if replace then let nexthyp = get_hyp_after id (named_context_of_val sign) in let sign,t,cl,sigma = clear_hyps2 env sigma (Id.Set.singleton id) sign t cl in - move_hyp false ([], LocalAssum (id,t),named_context_of_val sign) + move_hyp sigma false ([], LocalAssum (id,t),named_context_of_val sign) nexthyp, t,cl,sigma else diff --git a/proofs/logic.mli b/proofs/logic.mli index 0dba9ef1e..8c8a01cad 100644 --- a/proofs/logic.mli +++ b/proofs/logic.mli @@ -57,5 +57,5 @@ val catchable_exception : exn -> bool val convert_hyp : bool -> Environ.named_context_val -> evar_map -> Context.Named.Declaration.t -> Environ.named_context_val -val move_hyp_in_named_context : Id.t -> Id.t Misctypes.move_location -> +val move_hyp_in_named_context : Evd.evar_map -> Id.t -> Id.t Misctypes.move_location -> Environ.named_context_val -> Environ.named_context_val diff --git a/proofs/redexpr.ml b/proofs/redexpr.ml index 34443b93d..a442a5e63 100644 --- a/proofs/redexpr.ml +++ b/proofs/redexpr.ml @@ -25,7 +25,7 @@ open Misctypes (* call by value normalisation function using the virtual machine *) let cbv_vm env sigma c = let ctyp = Retyping.get_type_of env sigma c in - if Termops.occur_meta_or_existential c then + if Termops.occur_meta_or_existential sigma (EConstr.of_constr c) then error "vm_compute does not support existential variables."; Vnorm.cbv_vm env c ctyp diff --git a/tactics/auto.ml b/tactics/auto.ml index d4251555d..17fe7362d 100644 --- a/tactics/auto.ml +++ b/tactics/auto.ml @@ -302,7 +302,7 @@ let hintmap_of secvars hdc concl = match hdc with | None -> Hint_db.map_none ~secvars | Some hdc -> - if occur_existential concl then + if occur_existential Evd.empty (EConstr.of_constr concl) then (** FIXME *) Hint_db.map_existential ~secvars hdc concl else Hint_db.map_auto ~secvars hdc concl @@ -329,11 +329,12 @@ let rec trivial_fail_db dbg mod_delta db_list local_db = in Proofview.Goal.enter { enter = begin fun gl -> let concl = Tacmach.New.pf_nf_concl gl in + let sigma = Tacmach.New.project gl in let secvars = compute_secvars gl in Tacticals.New.tclFIRST ((dbg_assumption dbg)::intro_tac:: (List.map Tacticals.New.tclCOMPLETE - (trivial_resolve dbg mod_delta db_list local_db secvars concl))) + (trivial_resolve sigma dbg mod_delta db_list local_db secvars concl))) end } and my_find_search_nodelta db_list local_db secvars hdc concl = @@ -346,7 +347,7 @@ and my_find_search mod_delta = and my_find_search_delta db_list local_db secvars hdc concl = let f = hintmap_of secvars hdc concl in - if occur_existential concl then + if occur_existential Evd.empty (EConstr.of_constr concl) (** FIXME *) then List.map_append (fun db -> if Hint_db.use_dn db then @@ -402,10 +403,10 @@ and tac_of_hint dbg db_list local_db concl (flags, ({pat=p; code=t;poly=poly;db= in tclLOG dbg pr_hint (run_hint t tactic) -and trivial_resolve dbg mod_delta db_list local_db secvars cl = +and trivial_resolve sigma dbg mod_delta db_list local_db secvars cl = try let head = - try let hdconstr = decompose_app_bound cl in + try let hdconstr = decompose_app_bound sigma cl in Some hdconstr with Bound -> None in @@ -449,10 +450,10 @@ let h_trivial ?(debug=Off) lems l = gen_trivial ~debug lems l (* The classical Auto tactic *) (**************************************************************************) -let possible_resolve dbg mod_delta db_list local_db secvars cl = +let possible_resolve sigma dbg mod_delta db_list local_db secvars cl = try let head = - try let hdconstr = decompose_app_bound cl in + try let hdconstr = decompose_app_bound sigma cl in Some hdconstr with Bound -> None in @@ -488,12 +489,13 @@ let search d n mod_delta db_list local_db = (Tacticals.New.tclORELSE0 (intro_register d (search d n) local_db) ( Proofview.Goal.enter { enter = begin fun gl -> let concl = Tacmach.New.pf_nf_concl gl in + let sigma = Tacmach.New.project gl in let secvars = compute_secvars gl in let d' = incr_dbg d in Tacticals.New.tclFIRST (List.map (fun ntac -> Tacticals.New.tclTHEN ntac (search d' (n-1) local_db)) - (possible_resolve d mod_delta db_list local_db secvars concl)) + (possible_resolve sigma d mod_delta db_list local_db secvars concl)) end })) end [] in diff --git a/tactics/class_tactics.ml b/tactics/class_tactics.ml index a4243164e..fe7a09f77 100644 --- a/tactics/class_tactics.ml +++ b/tactics/class_tactics.ml @@ -279,9 +279,9 @@ let clenv_of_prods poly nprods (c, clenv) gl = let (c, _, _) = c in if poly || Int.equal nprods 0 then Some (None, clenv) else - let ty = Retyping.get_type_of (Proofview.Goal.env gl) - (Sigma.to_evar_map (Proofview.Goal.sigma gl)) c in - let diff = nb_prod ty - nprods in + let sigma = Tacmach.New.project gl in + let ty = Retyping.get_type_of (Proofview.Goal.env gl) sigma c in + let diff = nb_prod sigma (EConstr.of_constr ty) - nprods in if Pervasives.(>=) diff 0 then (* Was Some clenv... *) Some (Some diff, @@ -454,13 +454,13 @@ and e_my_find_search db_list local_db secvars hdc complete only_classes sigma co and e_trivial_resolve db_list local_db secvars only_classes sigma concl = try e_my_find_search db_list local_db secvars - (decompose_app_bound concl) true only_classes sigma concl + (decompose_app_bound sigma concl) true only_classes sigma concl with Bound | Not_found -> [] let e_possible_resolve db_list local_db secvars only_classes sigma concl = try e_my_find_search db_list local_db secvars - (decompose_app_bound concl) false only_classes sigma concl + (decompose_app_bound sigma concl) false only_classes sigma concl with Bound | Not_found -> [] let cut_of_hints h = @@ -666,7 +666,7 @@ module V85 = struct let needs_backtrack env evd oev concl = if Option.is_empty oev || is_Prop env evd concl then - occur_existential concl + occur_existential evd (EConstr.of_constr concl) else true let hints_tac hints sk fk {it = gl,info; sigma = s} = @@ -740,7 +740,7 @@ module V85 = struct let fk' = (fun e -> let do_backtrack = - if unique then occur_existential concl + if unique then occur_existential s' (EConstr.of_constr concl) else if info.unique then true else if List.is_empty gls' then needs_backtrack env s' info.is_evar concl @@ -975,7 +975,7 @@ module Search = struct NOT backtrack. *) let needs_backtrack env evd unique concl = if unique || is_Prop env evd concl then - occur_existential concl + occur_existential evd (EConstr.of_constr concl) else true let mark_unresolvables sigma goals = @@ -1486,16 +1486,17 @@ let _ = (** Take the head of the arity of a constr. Used in the partial application tactic. *) -let rec head_of_constr t = - let t = strip_outer_cast(collapse_appl t) in +let rec head_of_constr sigma t = + let t = strip_outer_cast sigma (EConstr.of_constr (collapse_appl sigma (EConstr.of_constr t))) in match kind_of_term t with - | Prod (_,_,c2) -> head_of_constr c2 - | LetIn (_,_,_,c2) -> head_of_constr c2 - | App (f,args) -> head_of_constr f + | Prod (_,_,c2) -> head_of_constr sigma c2 + | LetIn (_,_,_,c2) -> head_of_constr sigma c2 + | App (f,args) -> head_of_constr sigma f | _ -> t let head_of_constr h c = - let c = head_of_constr c in + Proofview.tclEVARMAP >>= fun sigma -> + let c = head_of_constr sigma c in letin_tac None (Name h) c None Locusops.allHyps let not_evar c = diff --git a/tactics/contradiction.ml b/tactics/contradiction.ml index 6b29f574c..fcbad4bf0 100644 --- a/tactics/contradiction.ml +++ b/tactics/contradiction.ml @@ -66,12 +66,12 @@ let contradiction_context = let id = NamedDecl.get_id d in let typ = nf_evar sigma (NamedDecl.get_type d) in let typ = whd_all env sigma typ in - if is_empty_type typ then + if is_empty_type sigma typ then simplest_elim (mkVar id) else match kind_of_term typ with - | Prod (na,t,u) when is_empty_type u -> + | Prod (na,t,u) when is_empty_type sigma u -> let is_unit_or_eq = - if use_negated_unit_or_eq_type () then match_with_unit_or_eq_type t + if use_negated_unit_or_eq_type () then match_with_unit_or_eq_type sigma t else None in Tacticals.New.tclORELSE (match is_unit_or_eq with @@ -105,7 +105,7 @@ let is_negation_of env sigma typ t = match kind_of_term (whd_all env sigma t) with | Prod (na,t,u) -> let u = nf_evar sigma u in - is_empty_type u && is_conv_leq env sigma typ t + is_empty_type sigma u && is_conv_leq env sigma typ t | _ -> false let contradiction_term (c,lbind as cl) = @@ -115,7 +115,7 @@ let contradiction_term (c,lbind as cl) = let type_of = Tacmach.New.pf_unsafe_type_of gl in let typ = type_of c in let _, ccl = splay_prod env sigma typ in - if is_empty_type ccl then + if is_empty_type sigma ccl then Tacticals.New.tclTHEN (elim false None cl None) (Tacticals.New.tclTRY assumption) diff --git a/tactics/eauto.ml b/tactics/eauto.ml index 10c975b8d..6250fef2d 100644 --- a/tactics/eauto.ml +++ b/tactics/eauto.ml @@ -32,7 +32,8 @@ let e_give_exact ?(flags=eauto_unif_flags) c = Proofview.Goal.enter { enter = begin fun gl -> let t1 = Tacmach.New.pf_unsafe_type_of gl c in let t2 = Tacmach.New.pf_concl (Proofview.Goal.assume gl) in - if occur_existential t1 || occur_existential t2 then + let sigma = Tacmach.New.project gl in + if occur_existential sigma (EConstr.of_constr t1) || occur_existential sigma (EConstr.of_constr t2) then Tacticals.New.tclTHEN (Clenvtac.unify ~flags t1) (exact_no_check c) else exact_check c end } @@ -123,7 +124,7 @@ let hintmap_of secvars hdc concl = match hdc with | None -> fun db -> Hint_db.map_none ~secvars db | Some hdc -> - if occur_existential concl then + if occur_existential Evd.empty (EConstr.of_constr concl) then (** FIXME *) (fun db -> Hint_db.map_existential ~secvars hdc concl db) else (fun db -> Hint_db.map_auto ~secvars hdc concl db) (* FIXME: should be (Hint_db.map_eauto hdc concl db) *) @@ -147,7 +148,7 @@ let rec e_trivial_fail_db db_list local_db = let tacl = registered_e_assumption :: (Tacticals.New.tclTHEN Tactics.intro next) :: - (List.map fst (e_trivial_resolve db_list local_db secvars (Tacmach.New.pf_nf_concl gl))) + (List.map fst (e_trivial_resolve (Tacmach.New.project gl) db_list local_db secvars (Tacmach.New.pf_nf_concl gl))) in Tacticals.New.tclFIRST (List.map Tacticals.New.tclCOMPLETE tacl) end } @@ -181,13 +182,13 @@ and e_my_find_search db_list local_db secvars hdc concl = in List.map tac_of_hint hintl -and e_trivial_resolve db_list local_db secvars gl = - let hd = try Some (decompose_app_bound gl) with Bound -> None in +and e_trivial_resolve sigma db_list local_db secvars gl = + let hd = try Some (decompose_app_bound sigma gl) with Bound -> None in try priority (e_my_find_search db_list local_db secvars hd gl) with Not_found -> [] -let e_possible_resolve db_list local_db secvars gl = - let hd = try Some (decompose_app_bound gl) with Bound -> None in +let e_possible_resolve sigma db_list local_db secvars gl = + let hd = try Some (decompose_app_bound sigma gl) with Bound -> None in try List.map (fun (b, (tac, pp)) -> (tac, b, pp)) (e_my_find_search db_list local_db secvars hd gl) with Not_found -> [] @@ -289,7 +290,7 @@ module SearchProblem = struct let l = let concl = Reductionops.nf_evar (project g)(pf_concl g) in filter_tactics s.tacres - (e_possible_resolve s.dblist (List.hd s.localdb) secvars concl) + (e_possible_resolve (project g) s.dblist (List.hd s.localdb) secvars concl) in List.map (fun (lgls, cost, pp) -> diff --git a/tactics/elim.ml b/tactics/elim.ml index 3f0c01a29..12d8e98c4 100644 --- a/tactics/elim.ml +++ b/tactics/elim.ml @@ -79,11 +79,12 @@ let up_to_delta = ref false (* true *) let general_decompose recognizer c = Proofview.Goal.enter { enter = begin fun gl -> let type_of = pf_unsafe_type_of gl in + let sigma = project gl in let typc = type_of c in tclTHENS (cut typc) [ tclTHEN (intro_using tmphyp_name) (onLastHypId - (ifOnHyp recognizer (general_decompose_aux recognizer) + (ifOnHyp (recognizer sigma) (general_decompose_aux (recognizer sigma)) (fun id -> clear [id]))); exact_no_check c ] end } @@ -102,17 +103,17 @@ let head_in indl t gl = let decompose_these c l = Proofview.Goal.enter { enter = begin fun gl -> let indl = List.map (fun x -> x, Univ.Instance.empty) l in - general_decompose (fun (_,t) -> head_in indl t gl) c + general_decompose (fun sigma (_,t) -> head_in indl t gl) c end } let decompose_and c = general_decompose - (fun (_,t) -> is_record t) + (fun sigma (_,t) -> is_record sigma t) c let decompose_or c = general_decompose - (fun (_,t) -> is_disjunction t) + (fun sigma (_,t) -> is_disjunction sigma t) c let h_decompose l c = decompose_these c l @@ -133,7 +134,7 @@ let induction_trailer abs_i abs_j bargs = (fun id -> Proofview.Goal.nf_enter { enter = begin fun gl -> let idty = pf_unsafe_type_of gl (mkVar id) in - let fvty = global_vars (pf_env gl) idty in + let fvty = global_vars (pf_env gl) (project gl) (EConstr.of_constr idty) in let possible_bring_hyps = (List.tl (nLastDecls gl (abs_j - abs_i))) @ bargs.Tacticals.assums in diff --git a/tactics/equality.ml b/tactics/equality.ml index 7c819edad..74f6dd44a 100644 --- a/tactics/equality.ml +++ b/tactics/equality.ml @@ -405,7 +405,7 @@ let leibniz_rewrite_ebindings_clause cls lft2rgt tac c t l with_evars frzevars d let isatomic = isProd (whd_zeta evd hdcncl) in let dep_fun = if isatomic then dependent else dependent_no_evar in let type_of_cls = type_of_clause cls gl in - let dep = dep_proof_ok && dep_fun c type_of_cls in + let dep = dep_proof_ok && dep_fun evd (EConstr.of_constr c) (EConstr.of_constr type_of_cls) in let Sigma ((elim, effs), sigma, p) = find_elim hdcncl lft2rgt dep cls (Some t) gl in let tac = Proofview.tclEFFECTS effs <*> @@ -442,7 +442,7 @@ let general_rewrite_ebindings_clause cls lft2rgt occs frzevars dep_proof_ok ?tac let env = Proofview.Goal.env gl in let ctype = get_type_of env sigma c in let rels, t = decompose_prod_assum (whd_betaiotazeta sigma ctype) in - match match_with_equality_type t with + match match_with_equality_type sigma t with | Some (hdcncl,args) -> (* Fast path: direct leibniz-like rewrite *) let lft2rgt = adjust_rewriting_direction args lft2rgt in leibniz_rewrite_ebindings_clause cls lft2rgt tac c (it_mkProd_or_LetIn t rels) @@ -455,9 +455,10 @@ let general_rewrite_ebindings_clause cls lft2rgt occs frzevars dep_proof_ok ?tac end begin function | (e, info) -> + Proofview.tclEVARMAP >>= fun sigma -> let env' = push_rel_context rels env in let rels',t' = splay_prod_assum env' sigma t in (* Search for underlying eq *) - match match_with_equality_type t' with + match match_with_equality_type sigma t' with | Some (hdcncl,args) -> let lft2rgt = adjust_rewriting_direction args lft2rgt in leibniz_rewrite_ebindings_clause cls lft2rgt tac c @@ -932,9 +933,10 @@ let rec build_discriminator env sigma dirn c = function let gen_absurdity id = Proofview.Goal.enter { enter = begin fun gl -> + let sigma = project gl in let hyp_typ = pf_get_hyp_typ id (Proofview.Goal.assume gl) in let hyp_typ = pf_nf_evar gl hyp_typ in - if is_empty_type hyp_typ + if is_empty_type sigma hyp_typ then simplest_elim (mkVar id) else @@ -973,7 +975,7 @@ let apply_on_clause (f,t) clause = let sigma = clause.evd in let f_clause = mk_clenv_from_env clause.env sigma None (f,t) in let argmv = - (match kind_of_term (last_arg f_clause.templval.Evd.rebus) with + (match kind_of_term (last_arg f_clause.evd (EConstr.of_constr f_clause.templval.Evd.rebus)) with | Meta mv -> mv | _ -> user_err (str "Ill-formed clause applicator.")) in clenv_fchain ~with_univs:false argmv f_clause clause @@ -1025,7 +1027,7 @@ let onNegatedEquality with_evars tac = let ccl = Proofview.Goal.concl gl in let env = Proofview.Goal.env gl in match kind_of_term (hnf_constr env sigma ccl) with - | Prod (_,t,u) when is_empty_type u -> + | Prod (_,t,u) when is_empty_type sigma u -> tclTHEN introf (onLastHypId (fun id -> onEquality with_evars tac (mkVar id,NoBindings))) @@ -1079,7 +1081,7 @@ let find_sigma_data env s = build_sigma_type () *) let make_tuple env sigma (rterm,rty) lind = - assert (dependent (mkRel lind) rty); + assert (not (EConstr.Vars.noccurn sigma lind (EConstr.of_constr rty))); let sigdata = find_sigma_data env (get_sort_of env sigma rty) in let sigma, a = type_of ~refresh:true env sigma (mkRel lind) in let na = Context.Rel.Declaration.get_name (lookup_rel lind env) in @@ -1101,9 +1103,9 @@ let make_tuple env sigma (rterm,rty) lind = normalization *) let minimal_free_rels env sigma (c,cty) = - let cty_rels = free_rels cty in + let cty_rels = free_rels sigma (EConstr.of_constr cty) in let cty' = simpl env sigma cty in - let rels' = free_rels cty' in + let rels' = free_rels sigma (EConstr.of_constr cty') in if Int.Set.subset cty_rels rels' then (cty,cty_rels) else @@ -1302,6 +1304,7 @@ let set_eq_dec_scheme_kind k = eq_dec_scheme_kind_name := (fun _ -> k) let inject_if_homogenous_dependent_pair ty = Proofview.Goal.nf_enter { enter = begin fun gl -> try + let sigma = Tacmach.New.project gl in let eq,u,(t,t1,t2) = find_this_eq_data_decompose gl ty in (* fetch the informations of the pair *) let ceq = Universes.constr_of_global Coqlib.glob_eq in @@ -1310,8 +1313,8 @@ let inject_if_homogenous_dependent_pair ty = (* check whether the equality deals with dep pairs or not *) let eqTypeDest = fst (decompose_app t) in if not (Globnames.is_global (sigTconstr()) eqTypeDest) then raise Exit; - let hd1,ar1 = decompose_app_vect t1 and - hd2,ar2 = decompose_app_vect t2 in + let hd1,ar1 = decompose_app_vect sigma (EConstr.of_constr t1) and + hd2,ar2 = decompose_app_vect sigma (EConstr.of_constr t2) in if not (Globnames.is_global (existTconstr()) hd1) then raise Exit; if not (Globnames.is_global (existTconstr()) hd2) then raise Exit; let ind,_ = try pf_apply find_mrectype gl ar1.(0) with Not_found -> raise Exit in @@ -1543,7 +1546,7 @@ let subst_tuple_term env sigma dep_pair1 dep_pair2 b = (* We build the expected goal *) let abst_B = List.fold_right - (fun (e,t) body -> lambda_create env (t,subst_term e body)) e1_list b in + (fun (e,t) body -> lambda_create env (t,subst_term sigma (EConstr.of_constr e) (EConstr.of_constr body))) e1_list b in let pred_body = beta_applist(abst_B,proj_list) in let body = mkApp (lambda_create env (typ,pred_body),[|dep_pair1|]) in let expected_goal = beta_applist (abst_B,List.map fst e2_list) in @@ -1674,8 +1677,8 @@ let is_eq_x gl x d = in let c = pf_nf_evar gl (NamedDecl.get_type d) in let (_,lhs,rhs) = pi3 (find_eq_data_decompose gl c) in - if (is_var x lhs) && not (local_occur_var x rhs) then raise (FoundHyp (id,rhs,true)); - if (is_var x rhs) && not (local_occur_var x lhs) then raise (FoundHyp (id,lhs,false)) + if (is_var x lhs) && not (local_occur_var (project gl) x (EConstr.of_constr rhs)) then raise (FoundHyp (id,rhs,true)); + if (is_var x rhs) && not (local_occur_var (project gl) x (EConstr.of_constr lhs)) then raise (FoundHyp (id,lhs,false)) with Constr_matching.PatternMatchingFailure -> () @@ -1685,6 +1688,7 @@ let is_eq_x gl x d = let subst_one dep_proof_ok x (hyp,rhs,dir) = Proofview.Goal.enter { enter = begin fun gl -> let env = Proofview.Goal.env gl in + let sigma = Tacmach.New.project gl in let hyps = Proofview.Goal.hyps (Proofview.Goal.assume gl) in let concl = Proofview.Goal.concl (Proofview.Goal.assume gl) in (* The set of hypotheses using x *) @@ -1692,7 +1696,7 @@ let subst_one dep_proof_ok x (hyp,rhs,dir) = List.rev (pi3 (List.fold_right (fun dcl (dest,deps,allhyps) -> let id = NamedDecl.get_id dcl in if not (Id.equal id hyp) - && List.exists (fun y -> occur_var_in_decl env y dcl) deps + && List.exists (fun y -> occur_var_in_decl env sigma y dcl) deps then let id_dest = if !regular_subst_tactic then dest else MoveLast in (dest,id::deps,(id_dest,id)::allhyps) @@ -1701,7 +1705,7 @@ let subst_one dep_proof_ok x (hyp,rhs,dir) = hyps (MoveBefore x,[x],[]))) in (* In practice, no dep hyps before x, so MoveBefore x is good enough *) (* Decides if x appears in conclusion *) - let depconcl = occur_var env x concl in + let depconcl = occur_var env sigma x (EConstr.of_constr concl) in let need_rewrite = not (List.is_empty dephyps) || depconcl in tclTHENLIST ((if need_rewrite then @@ -1787,6 +1791,7 @@ let subst_all ?(flags=default_subst_tactic_flags ()) () = let process hyp = Proofview.Goal.enter { enter = begin fun gl -> let gl = Proofview.Goal.assume gl in + let sigma = project gl in let env = Proofview.Goal.env gl in let find_eq_data_decompose = find_eq_data_decompose gl in let c = pf_get_hyp hyp gl |> NamedDecl.get_type in @@ -1794,9 +1799,9 @@ let subst_all ?(flags=default_subst_tactic_flags ()) () = (* J.F.: added to prevent failure on goal containing x=x as an hyp *) if Term.eq_constr x y then Proofview.tclUNIT () else match kind_of_term x, kind_of_term y with - | Var x', _ when not (occur_term x y) && not (is_evaluable env (EvalVarRef x')) -> + | Var x', _ when not (occur_term sigma (EConstr.of_constr x) (EConstr.of_constr y)) && not (is_evaluable env (EvalVarRef x')) -> subst_one flags.rewrite_dependent_proof x' (hyp,y,true) - | _, Var y' when not (occur_term y x) && not (is_evaluable env (EvalVarRef y')) -> + | _, Var y' when not (occur_term sigma (EConstr.of_constr y) (EConstr.of_constr x)) && not (is_evaluable env (EvalVarRef y')) -> subst_one flags.rewrite_dependent_proof y' (hyp,x,false) | _ -> Proofview.tclUNIT () diff --git a/tactics/hints.ml b/tactics/hints.ml index 9fa49264f..55bf5f29e 100644 --- a/tactics/hints.ml +++ b/tactics/hints.ml @@ -45,8 +45,8 @@ type debug = Debug | Info | Off exception Bound -let head_constr_bound t = - let t = strip_outer_cast t in +let head_constr_bound sigma t = + let t = strip_outer_cast sigma (EConstr.of_constr t) in let _,ccl = decompose_prod_assum t in let hd,args = decompose_app ccl in match kind_of_term hd with @@ -54,13 +54,13 @@ let head_constr_bound t = | Proj (p, _) -> mkConst (Projection.constant p) | _ -> raise Bound -let head_constr c = - try head_constr_bound c with Bound -> error "Bound head variable." +let head_constr sigma c = + try head_constr_bound sigma c with Bound -> error "Bound head variable." -let decompose_app_bound t = - let t = strip_outer_cast t in +let decompose_app_bound sigma t = + let t = strip_outer_cast sigma (EConstr.of_constr t) in let _,ccl = decompose_prod_assum t in - let hd,args = decompose_app_vect ccl in + let hd,args = decompose_app_vect sigma (EConstr.of_constr ccl) in match kind_of_term hd with | Const (c,u) -> ConstRef c, args | Ind (i,u) -> IndRef i, args @@ -505,7 +505,7 @@ struct let match_mode m arg = match m with - | ModeInput -> not (occur_existential arg) + | ModeInput -> not (occur_existential Evd.empty (EConstr.of_constr arg)) (** FIXME *) | ModeNoHeadEvar -> Evarutil.(try ignore(head_evar arg); false with NoHeadEvar -> true) @@ -742,7 +742,7 @@ let secvars_of_global env gr = let make_exact_entry env sigma pri poly ?(name=PathAny) (c, cty, ctx) = let secvars = secvars_of_constr env c in - let cty = strip_outer_cast cty in + let cty = strip_outer_cast sigma (EConstr.of_constr cty) in match kind_of_term cty with | Prod _ -> failwith "make_exact_entry" | _ -> @@ -911,7 +911,7 @@ let make_trivial env sigma poly ?(name=PathAny) r = let c,ctx = fresh_global_or_constr env sigma poly r in let sigma = Evd.merge_context_set univ_flexible sigma ctx in let t = hnf_constr env sigma (unsafe_type_of env sigma c) in - let hd = head_of_constr_reference (head_constr t) in + let hd = head_of_constr_reference (head_constr sigma t) in let ce = mk_clenv_from_env env sigma None (c,t) in (Some hd, { pri=1; poly = poly; @@ -1013,7 +1013,7 @@ let subst_autohint (subst, obj) = let subst_key gr = let (lab'', elab') = subst_global subst gr in let gr' = - (try head_of_constr_reference (head_constr_bound elab') + (try head_of_constr_reference (head_constr_bound Evd.empty (** FIXME *) elab') with Bound -> lab'') in if gr' == gr then gr else gr' in @@ -1190,17 +1190,17 @@ let prepare_hint check (poly,local) env init (sigma,c) = thing make_resolves will do is to re-instantiate the products *) let sigma, subst = Evd.nf_univ_variables sigma in let c = Vars.subst_univs_constr subst (Evarutil.nf_evar sigma c) in - let c = drop_extra_implicit_args c in - let vars = ref (collect_vars c) in + let c = drop_extra_implicit_args sigma (EConstr.of_constr c) in + let vars = ref (collect_vars sigma (EConstr.of_constr c)) in let subst = ref [] in let rec find_next_evar c = match kind_of_term c with | Evar (evk,args as ev) -> (* We skip the test whether args is the identity or not *) let t = Evarutil.nf_evar sigma (existential_type sigma ev) in - let t = List.fold_right (fun (e,id) c -> replace_term e id c) !subst t in + let t = List.fold_right (fun (e,id) c -> replace_term sigma (EConstr.of_constr e) (EConstr.of_constr id) (EConstr.of_constr c)) !subst t in if not (closed0 c) then error "Hints with holes dependent on a bound variable not supported."; - if occur_existential t then + if occur_existential sigma (EConstr.of_constr t) then (* Not clever enough to construct dependency graph of evars *) error "Not clever enough to deal with evars dependent in other evars."; raise (Found (c,t)) @@ -1211,7 +1211,7 @@ let prepare_hint check (poly,local) env init (sigma,c) = let id = next_ident_away_from default_prepare_hint_ident (fun id -> Id.Set.mem id !vars) in vars := Id.Set.add id !vars; subst := (evar,mkVar id)::!subst; - mkNamedLambda id t (iter (replace_term evar (mkVar id) c)) in + mkNamedLambda id t (iter (replace_term sigma (EConstr.of_constr evar) (EConstr.mkVar id) (EConstr.of_constr c))) in let c' = iter c in if check then Pretyping.check_evars (Global.env()) Evd.empty sigma c'; let diff = Univ.ContextSet.diff (Evd.universe_context_set sigma) (Evd.universe_context_set init) in @@ -1394,13 +1394,13 @@ let pr_hint_ref ref = pr_hint_list_for_head ref (* Print all hints associated to head id in any database *) -let pr_hint_term cl = +let pr_hint_term sigma cl = try let dbs = current_db () in let valid_dbs = let fn = try - let hdc = decompose_app_bound cl in - if occur_existential cl then + let hdc = decompose_app_bound sigma cl in + if occur_existential sigma (EConstr.of_constr cl) then Hint_db.map_existential ~secvars:Id.Pred.full hdc cl else Hint_db.map_auto ~secvars:Id.Pred.full hdc cl with Bound -> Hint_db.map_none ~secvars:Id.Pred.full @@ -1423,7 +1423,7 @@ let pr_applicable_hint () = match glss.Evd.it with | [] -> CErrors.error "No focused goal." | g::_ -> - pr_hint_term (Goal.V82.concl glss.Evd.sigma g) + pr_hint_term glss.Evd.sigma (Goal.V82.concl glss.Evd.sigma g) let pp_hint_mode = function | ModeInput -> str"+" diff --git a/tactics/hints.mli b/tactics/hints.mli index edc65c407..c0eb2c3b8 100644 --- a/tactics/hints.mli +++ b/tactics/hints.mli @@ -24,7 +24,7 @@ open Vernacexpr exception Bound -val decompose_app_bound : constr -> global_reference * constr array +val decompose_app_bound : evar_map -> constr -> global_reference * constr array type debug = Debug | Info | Off diff --git a/tactics/hipattern.ml b/tactics/hipattern.ml index 27af7200b..847ecf4b0 100644 --- a/tactics/hipattern.ml +++ b/tactics/hipattern.ml @@ -31,9 +31,9 @@ module RelDecl = Context.Rel.Declaration -- Eduardo (6/8/97). *) -type 'a matching_function = constr -> 'a option +type 'a matching_function = Evd.evar_map -> constr -> 'a option -type testing_function = constr -> bool +type testing_function = Evd.evar_map -> constr -> bool let mkmeta n = Nameops.make_ident "X" (Some n) let meta1 = mkmeta 1 @@ -43,7 +43,7 @@ let meta4 = mkmeta 4 let op2bool = function Some _ -> true | None -> false -let match_with_non_recursive_type t = +let match_with_non_recursive_type sigma t = match kind_of_term t with | App _ -> let (hdapp,args) = decompose_app t in @@ -56,21 +56,21 @@ let match_with_non_recursive_type t = | _ -> None) | _ -> None -let is_non_recursive_type t = op2bool (match_with_non_recursive_type t) +let is_non_recursive_type sigma t = op2bool (match_with_non_recursive_type sigma t) (* Test dependencies *) (* NB: we consider also the let-in case in the following function, since they may appear in types of inductive constructors (see #2629) *) -let rec has_nodep_prod_after n c = +let rec has_nodep_prod_after n sigma c = match kind_of_term c with | Prod (_,_,b) | LetIn (_,_,_,b) -> - ( n>0 || not (dependent (mkRel 1) b)) - && (has_nodep_prod_after (n-1) b) + ( n>0 || EConstr.Vars.noccurn sigma 1 (EConstr.of_constr b)) + && (has_nodep_prod_after (n-1) sigma b) | _ -> true -let has_nodep_prod = has_nodep_prod_after 0 +let has_nodep_prod sigma c = has_nodep_prod_after 0 sigma c (* A general conjunctive type is a non-recursive with-no-indices inductive type with only one constructor and no dependencies between argument; @@ -87,7 +87,7 @@ let is_lax_conjunction = function | Some false -> true | _ -> false -let match_with_one_constructor style onlybinary allow_rec t = +let match_with_one_constructor sigma style onlybinary allow_rec t = let (hdapp,args) = decompose_app t in let res = match kind_of_term hdapp with | Ind ind -> @@ -112,7 +112,7 @@ let match_with_one_constructor style onlybinary allow_rec t = else let ctyp = prod_applist mip.mind_nf_lc.(0) args in let cargs = List.map RelDecl.get_type (prod_assum ctyp) in - if not (is_lax_conjunction style) || has_nodep_prod ctyp then + if not (is_lax_conjunction style) || has_nodep_prod sigma ctyp then (* Record or non strict conjunction *) Some (hdapp,List.rev cargs) else @@ -125,28 +125,28 @@ let match_with_one_constructor style onlybinary allow_rec t = | Some (hdapp, [_; _]) -> res | _ -> None -let match_with_conjunction ?(strict=false) ?(onlybinary=false) t = - match_with_one_constructor (Some strict) onlybinary false t +let match_with_conjunction ?(strict=false) ?(onlybinary=false) sigma t = + match_with_one_constructor sigma (Some strict) onlybinary false t -let match_with_record t = - match_with_one_constructor None false false t +let match_with_record sigma t = + match_with_one_constructor sigma None false false t -let is_conjunction ?(strict=false) ?(onlybinary=false) t = - op2bool (match_with_conjunction ~strict ~onlybinary t) +let is_conjunction ?(strict=false) ?(onlybinary=false) sigma t = + op2bool (match_with_conjunction sigma ~strict ~onlybinary t) -let is_record t = - op2bool (match_with_record t) +let is_record sigma t = + op2bool (match_with_record sigma t) -let match_with_tuple t = - let t = match_with_one_constructor None false true t in +let match_with_tuple sigma t = + let t = match_with_one_constructor sigma None false true t in Option.map (fun (hd,l) -> let ind = destInd hd in let (mib,mip) = Global.lookup_pinductive ind in let isrec = mis_is_recursive (fst ind,mib,mip) in (hd,l,isrec)) t -let is_tuple t = - op2bool (match_with_tuple t) +let is_tuple sigma t = + op2bool (match_with_tuple sigma t) (* A general disjunction type is a non-recursive with-no-indices inductive type with of which all constructors have a single argument; @@ -159,7 +159,7 @@ let test_strict_disjunction n lc = | [LocalAssum (_,c)] -> isRel c && Int.equal (destRel c) (n - i) | _ -> false) 0 lc -let match_with_disjunction ?(strict=false) ?(onlybinary=false) t = +let match_with_disjunction ?(strict=false) ?(onlybinary=false) sigma t = let (hdapp,args) = decompose_app t in let res = match kind_of_term hdapp with | Ind (ind,u) -> @@ -187,13 +187,13 @@ let match_with_disjunction ?(strict=false) ?(onlybinary=false) t = | Some (hdapp,[_; _]) -> res | _ -> None -let is_disjunction ?(strict=false) ?(onlybinary=false) t = - op2bool (match_with_disjunction ~strict ~onlybinary t) +let is_disjunction ?(strict=false) ?(onlybinary=false) sigma t = + op2bool (match_with_disjunction ~strict ~onlybinary sigma t) (* An empty type is an inductive type, possible with indices, that has no constructors *) -let match_with_empty_type t = +let match_with_empty_type sigma t = let (hdapp,args) = decompose_app t in match (kind_of_term hdapp) with | Ind ind -> @@ -202,33 +202,33 @@ let match_with_empty_type t = if Int.equal nconstr 0 then Some hdapp else None | _ -> None -let is_empty_type t = op2bool (match_with_empty_type t) +let is_empty_type sigma t = op2bool (match_with_empty_type sigma t) (* This filters inductive types with one constructor with no arguments; Parameters and indices are allowed *) -let match_with_unit_or_eq_type t = +let match_with_unit_or_eq_type sigma t = let (hdapp,args) = decompose_app t in match (kind_of_term hdapp) with | Ind ind -> let (mib,mip) = Global.lookup_pinductive ind in let constr_types = mip.mind_nf_lc in let nconstr = Array.length mip.mind_consnames in - let zero_args c = Int.equal (nb_prod c) mib.mind_nparams in + let zero_args c = Int.equal (nb_prod sigma (EConstr.of_constr c)) mib.mind_nparams in if Int.equal nconstr 1 && zero_args constr_types.(0) then Some hdapp else None | _ -> None -let is_unit_or_eq_type t = op2bool (match_with_unit_or_eq_type t) +let is_unit_or_eq_type sigma t = op2bool (match_with_unit_or_eq_type sigma t) (* A unit type is an inductive type with no indices but possibly (useless) parameters, and that has no arguments in its unique constructor *) -let is_unit_type t = - match match_with_conjunction t with +let is_unit_type sigma t = + match match_with_conjunction sigma t with | Some (_,[]) -> true | _ -> false @@ -318,13 +318,13 @@ let is_inductive_equality ind = let nconstr = Array.length mip.mind_consnames in Int.equal nconstr 1 && Int.equal (constructor_nrealargs (ind,1)) 0 -let match_with_equality_type t = +let match_with_equality_type sigma t = let (hdapp,args) = decompose_app t in match (kind_of_term hdapp) with | Ind (ind,_) when is_inductive_equality ind -> Some (hdapp,args) | _ -> None -let is_equality_type t = op2bool (match_with_equality_type t) +let is_equality_type sigma t = op2bool (match_with_equality_type sigma t) (* Arrows/Implication/Negation *) @@ -338,37 +338,37 @@ let match_arrow_pattern t = assert (Id.equal m1 meta1 && Id.equal m2 meta2); (arg, mind) | _ -> anomaly (Pp.str "Incorrect pattern matching") -let match_with_imp_term c= +let match_with_imp_term sigma c = match kind_of_term c with - | Prod (_,a,b) when not (dependent (mkRel 1) b) ->Some (a,b) + | Prod (_,a,b) when EConstr.Vars.noccurn sigma 1 (EConstr.of_constr b) -> Some (a,b) | _ -> None -let is_imp_term c = op2bool (match_with_imp_term c) +let is_imp_term sigma c = op2bool (match_with_imp_term sigma c) -let match_with_nottype t = +let match_with_nottype sigma t = try let (arg,mind) = match_arrow_pattern t in - if is_empty_type mind then Some (mind,arg) else None + if is_empty_type sigma mind then Some (mind,arg) else None with PatternMatchingFailure -> None -let is_nottype t = op2bool (match_with_nottype t) +let is_nottype sigma t = op2bool (match_with_nottype sigma t) (* Forall *) -let match_with_forall_term c= +let match_with_forall_term sigma c= match kind_of_term c with | Prod (nam,a,b) -> Some (nam,a,b) | _ -> None -let is_forall_term c = op2bool (match_with_forall_term c) +let is_forall_term sigma c = op2bool (match_with_forall_term sigma c) -let match_with_nodep_ind t = +let match_with_nodep_ind sigma t = let (hdapp,args) = decompose_app t in match (kind_of_term hdapp) with | Ind ind -> let (mib,mip) = Global.lookup_pinductive ind in if Array.length (mib.mind_packets)>1 then None else - let nodep_constr = has_nodep_prod_after mib.mind_nparams in + let nodep_constr = has_nodep_prod_after mib.mind_nparams sigma in if Array.for_all nodep_constr mip.mind_nf_lc then let params= if Int.equal mip.mind_nrealargs 0 then args else @@ -378,9 +378,9 @@ let match_with_nodep_ind t = None | _ -> None -let is_nodep_ind t=op2bool (match_with_nodep_ind t) +let is_nodep_ind sigma t = op2bool (match_with_nodep_ind sigma t) -let match_with_sigma_type t= +let match_with_sigma_type sigma t = let (hdapp,args) = decompose_app t in match (kind_of_term hdapp) with | Ind ind -> @@ -388,14 +388,14 @@ let match_with_sigma_type t= if Int.equal (Array.length (mib.mind_packets)) 1 && (Int.equal mip.mind_nrealargs 0) && (Int.equal (Array.length mip.mind_consnames)1) && - has_nodep_prod_after (mib.mind_nparams+1) mip.mind_nf_lc.(0) then + has_nodep_prod_after (mib.mind_nparams+1) sigma mip.mind_nf_lc.(0) then (*allowing only 1 existential*) Some (hdapp,args) else None | _ -> None -let is_sigma_type t=op2bool (match_with_sigma_type t) +let is_sigma_type sigma t = op2bool (match_with_sigma_type sigma t) (***** Destructing patterns bound to some theory *) diff --git a/tactics/hipattern.mli b/tactics/hipattern.mli index 7cc41f1b9..8a453bf31 100644 --- a/tactics/hipattern.mli +++ b/tactics/hipattern.mli @@ -40,8 +40,8 @@ open Coqlib also work on ad-hoc disjunctions introduced by the user. (Eduardo, 6/8/97). *) -type 'a matching_function = constr -> 'a option -type testing_function = constr -> bool +type 'a matching_function = Evd.evar_map -> constr -> 'a option +type testing_function = Evd.evar_map -> constr -> bool val match_with_non_recursive_type : (constr * constr list) matching_function val is_non_recursive_type : testing_function diff --git a/tactics/inv.ml b/tactics/inv.ml index e7d8249e4..d1d6178da 100644 --- a/tactics/inv.ml +++ b/tactics/inv.ml @@ -32,8 +32,9 @@ module NamedDecl = Context.Named.Declaration let var_occurs_in_pf gl id = let env = Proofview.Goal.env gl in - occur_var env id (Proofview.Goal.concl gl) || - List.exists (occur_var_in_decl env id) (Proofview.Goal.hyps gl) + let sigma = project gl in + occur_var env sigma id (EConstr.of_constr (Proofview.Goal.concl gl)) || + List.exists (occur_var_in_decl env sigma id) (Proofview.Goal.hyps gl) (* [make_inv_predicate (ity,args) C] @@ -75,7 +76,7 @@ let make_inv_predicate env evd indf realargs id status concl = let hyps_arity,_ = get_arity env indf in (hyps_arity,concl) | Dep dflt_concl -> - if not (occur_var env id concl) then + if not (occur_var env !evd id (EConstr.of_constr concl)) then user_err ~hdr:"make_inv_predicate" (str "Current goal does not depend on " ++ pr_id id ++ str"."); (* We abstract the conclusion of goal with respect to @@ -183,7 +184,7 @@ let dependent_hyps env id idlist gl = | d::l -> (* Update the type of id1: it may have been subject to rewriting *) let d = pf_get_hyp (NamedDecl.get_id d) gl in - if occur_var_in_decl env id d + if occur_var_in_decl env (project gl) id d then d :: dep_rec l else dep_rec l in @@ -448,7 +449,7 @@ let raw_inversion inv_kind id status names = make_inv_predicate env evdref indf realargs id status concl in let sigma = !evdref in let (cut_concl,case_tac) = - if status != NoDep && (dependent c concl) then + if status != NoDep && (dependent sigma (EConstr.of_constr c) (EConstr.of_constr concl)) then Reduction.beta_appvect elim_predicate (Array.of_list (realargs@[c])), case_then_using else @@ -514,12 +515,14 @@ let invIn k names ids id = Proofview.Goal.nf_enter { enter = begin fun gl -> let hyps = List.map (fun id -> pf_get_hyp id gl) ids in let concl = Proofview.Goal.concl gl in - let nb_prod_init = nb_prod concl in + let sigma = project gl in + let nb_prod_init = nb_prod sigma (EConstr.of_constr concl) in let intros_replace_ids = Proofview.Goal.enter { enter = begin fun gl -> let concl = pf_nf_concl gl in + let sigma = project gl in let nb_of_new_hyp = - nb_prod concl - (List.length hyps + nb_prod_init) + nb_prod sigma (EConstr.of_constr concl) - (List.length hyps + nb_prod_init) in if nb_of_new_hyp < 1 then intros_replacing ids diff --git a/tactics/leminv.ml b/tactics/leminv.ml index 10fc5076c..46f1f7c8d 100644 --- a/tactics/leminv.ml +++ b/tactics/leminv.ml @@ -154,7 +154,7 @@ let compute_first_inversion_scheme env sigma ind sort dep_option = pty,goal else let i = mkAppliedInd ind in - let ivars = global_vars env i in + let ivars = global_vars env sigma (EConstr.of_constr i) in let revargs,ownsign = fold_named_context (fun env d (revargs,hyps) -> @@ -192,7 +192,7 @@ let inversion_scheme env sigma t sort dep_option inv_op = in assert (List.subset - (global_vars env invGoal) + (global_vars env sigma (EConstr.of_constr invGoal)) (ids_of_named_context (named_context invEnv))); (* user_err ~hdr:"lemma_inversion" @@ -277,7 +277,8 @@ let lemInvIn id c ids = let hyps = List.map (fun id -> pf_get_hyp id gl) ids in let intros_replace_ids = let concl = Proofview.Goal.concl gl in - let nb_of_new_hyp = nb_prod concl - List.length ids in + let sigma = project gl in + let nb_of_new_hyp = nb_prod sigma (EConstr.of_constr concl) - List.length ids in if nb_of_new_hyp < 1 then intros_replacing ids else diff --git a/tactics/tacticals.ml b/tactics/tacticals.ml index 93c04e373..676b23d09 100644 --- a/tactics/tacticals.ml +++ b/tactics/tacticals.ml @@ -630,7 +630,7 @@ module New = struct (* applying elimination_scheme just a little modified *) let elimclause = Tacmach.New.of_old (fun gls -> mk_clenv_from gls (elim,Tacmach.New.pf_unsafe_type_of gl elim)) gl in let indmv = - match kind_of_term (last_arg elimclause.templval.Evd.rebus) with + match kind_of_term (last_arg elimclause.evd (EConstr.of_constr elimclause.templval.Evd.rebus)) with | Meta mv -> mv | _ -> anomaly (str"elimination") in diff --git a/tactics/tactics.ml b/tactics/tactics.ml index e17bbfcb0..15dd1a97c 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -323,10 +323,11 @@ let apply_clear_request clear_flag dft c = let move_hyp id dest = Proofview.Goal.enter { enter = begin fun gl -> let env = Proofview.Goal.env gl in + let sigma = Tacmach.New.project gl in let ty = Proofview.Goal.raw_concl gl in let store = Proofview.Goal.extra gl in let sign = named_context_val env in - let sign' = move_hyp_in_named_context id dest sign in + let sign' = move_hyp_in_named_context sigma id dest sign in let env = reset_with_named_context sign' env in Refine.refine ~unsafe:true { run = begin fun sigma -> Evarutil.new_evar env sigma ~principal:true ~store ty @@ -497,7 +498,7 @@ fun env sigma p -> function let Sigma (rem, sigma, r) = mk_holes env sigma (p +> q) rem in Sigma (arg :: rem, sigma, r) -let rec check_mutind env sigma k cl = match kind_of_term (strip_outer_cast cl) with +let rec check_mutind env sigma k cl = match kind_of_term (strip_outer_cast sigma (EConstr.of_constr cl)) with | Prod (na, c1, b) -> if Int.equal k 1 then try @@ -936,13 +937,14 @@ let build_intro_tac id dest tac = match dest with let rec intro_then_gen name_flag move_flag force_flag dep_flag tac = let open Context.Rel.Declaration in Proofview.Goal.enter { enter = begin fun gl -> + let sigma = Tacmach.New.project gl in let concl = Proofview.Goal.concl (Proofview.Goal.assume gl) in let concl = nf_evar (Tacmach.New.project gl) concl in match kind_of_term concl with - | Prod (name,t,u) when not dep_flag || (dependent (mkRel 1) u) -> + | Prod (name,t,u) when not dep_flag || not (EConstr.Vars.noccurn sigma 1 (EConstr.of_constr u)) -> let name = find_name false (LocalAssum (name,t)) name_flag gl in build_intro_tac name move_flag tac - | LetIn (name,b,t,u) when not dep_flag || (dependent (mkRel 1) u) -> + | LetIn (name,b,t,u) when not dep_flag || not (EConstr.Vars.noccurn sigma 1 (EConstr.of_constr u)) -> let name = find_name false (LocalDef (name,b,t)) name_flag gl in build_intro_tac name move_flag tac | _ -> @@ -1285,7 +1287,7 @@ let clenv_refine_in ?(sidecond_first=false) with_evars ?(with_classes=true) in let new_hyp_typ = clenv_type clenv in if not with_evars then check_unresolved_evars_of_metas sigma0 clenv; - if not with_evars && occur_meta new_hyp_typ then + if not with_evars && occur_meta clenv.evd (EConstr.of_constr new_hyp_typ) then error_uninstantiated_metas new_hyp_typ clenv; let new_hyp_prf = clenv_value clenv in let exact_tac = Proofview.V82.tactic (Tacmach.refine_no_check new_hyp_prf) in @@ -1440,7 +1442,7 @@ let general_case_analysis_in_context with_evars clear_flag (c,lbindc) = let (mind,_) = reduce_to_quantified_ind env (Sigma.to_evar_map sigma) t in let sort = Tacticals.New.elimination_sort_of_goal gl in let Sigma (elim, sigma, p) = - if occur_term c concl then + if occur_term (Sigma.to_evar_map sigma) (EConstr.of_constr c) (EConstr.of_constr concl) then build_case_analysis_scheme env sigma mind true sort else build_case_analysis_scheme_default env sigma mind sort in @@ -1624,7 +1626,7 @@ let descend_in_conjunctions avoid tac (err, info) c = let t = Retyping.get_type_of env sigma c in let ((ind,u),t) = reduce_to_quantified_ind env sigma t in let sign,ccl = decompose_prod_assum t in - match match_with_tuple ccl with + match match_with_tuple sigma ccl with | Some (_,_,isrec) -> let n = (constructors_nrealargs ind).(0) in let sort = Tacticals.New.elimination_sort_of_goal gl in @@ -1689,12 +1691,13 @@ let tclORELSEOPT t k = let general_apply with_delta with_destruct with_evars clear_flag (loc,(c,lbind)) = Proofview.Goal.nf_enter { enter = begin fun gl -> let concl = Proofview.Goal.concl gl in + let sigma = Tacmach.New.project gl in let flags = if with_delta then default_unify_flags () else default_no_delta_unify_flags () in (* The actual type of the theorem. It will be matched against the goal. If this fails, then the head constant will be unfolded step by step. *) - let concl_nprod = nb_prod_modulo_zeta concl in + let concl_nprod = nb_prod_modulo_zeta sigma (EConstr.of_constr concl) in let rec try_main_apply with_destruct c = Proofview.Goal.enter { enter = begin fun gl -> let env = Proofview.Goal.env gl in @@ -1703,7 +1706,7 @@ let general_apply with_delta with_destruct with_evars clear_flag (loc,(c,lbind)) let thm_ty0 = nf_betaiota sigma (Retyping.get_type_of env sigma c) in let try_apply thm_ty nprod = try - let n = nb_prod_modulo_zeta thm_ty - nprod in + let n = nb_prod_modulo_zeta sigma (EConstr.of_constr thm_ty) - nprod in if n<0 then error "Applied theorem has not enough premisses."; let clause = make_clenv_binding_apply env sigma (Some n) (c,thm_ty) lbind in Clenvtac.res_pf clause ~with_evars ~flags @@ -1901,8 +1904,9 @@ let apply_in_delayed_once sidecond_first with_delta with_destruct with_evars nam let cut_and_apply c = Proofview.Goal.nf_enter { enter = begin fun gl -> + let sigma = Tacmach.New.project gl in match kind_of_term (Tacmach.New.pf_hnf_constr gl (Tacmach.New.pf_unsafe_type_of gl c)) with - | Prod (_,c1,c2) when not (dependent (mkRel 1) c2) -> + | Prod (_,c1,c2) when EConstr.Vars.noccurn sigma 1 (EConstr.of_constr c2) -> let concl = Proofview.Goal.concl gl in let env = Tacmach.New.pf_env gl in Refine.refine { run = begin fun sigma -> @@ -2049,7 +2053,7 @@ let clear_body ids = (** Do no recheck hypotheses that do not depend *) let sigma = if not seen then sigma - else if List.exists (fun id -> occur_var_in_decl env id decl) ids then + else if List.exists (fun id -> occur_var_in_decl env sigma id decl) ids then check_decl env sigma decl else sigma in @@ -2058,7 +2062,7 @@ let clear_body ids = in let (env, sigma, _) = List.fold_left check (base_env, sigma, false) (List.rev ctx) in let sigma = - if List.exists (fun id -> occur_var env id concl) ids then + if List.exists (fun id -> occur_var env sigma id (EConstr.of_constr concl)) ids then check_is_type env sigma concl else sigma in @@ -2096,12 +2100,13 @@ let keep hyps = Proofview.Goal.nf_enter { enter = begin fun gl -> Proofview.tclENV >>= fun env -> let ccl = Proofview.Goal.concl gl in + let sigma = Tacmach.New.project gl in let cl,_ = fold_named_context_reverse (fun (clear,keep) decl -> let hyp = NamedDecl.get_id decl in if Id.List.mem hyp hyps - || List.exists (occur_var_in_decl env hyp) keep - || occur_var env hyp ccl + || List.exists (occur_var_in_decl env sigma hyp) keep + || occur_var env sigma hyp (EConstr.of_constr ccl) then (clear,decl::keep) else (hyp::clear,keep)) ~init:([],[]) (Proofview.Goal.env gl) @@ -2310,15 +2315,16 @@ let rewrite_hyp_then assert_style with_evars thin l2r id tac = List.filter (fun (_,id) -> not (Id.equal id id')) thin in Proofview.Goal.enter { enter = begin fun gl -> let env = Proofview.Goal.env gl in + let sigma = Tacmach.New.project gl in let type_of = Tacmach.New.pf_unsafe_type_of gl in let whd_all = Tacmach.New.pf_apply whd_all gl in let t = whd_all (type_of (mkVar id)) in - let eqtac, thin = match match_with_equality_type t with + let eqtac, thin = match match_with_equality_type sigma t with | Some (hdcncl,[_;lhs;rhs]) -> - if l2r && isVar lhs && not (occur_var env (destVar lhs) rhs) then + if l2r && isVar lhs && not (occur_var env sigma (destVar lhs) (EConstr.of_constr rhs)) then let id' = destVar lhs in subst_on l2r id' rhs, early_clear id' thin - else if not l2r && isVar rhs && not (occur_var env (destVar rhs) lhs) then + else if not l2r && isVar rhs && not (occur_var env sigma (destVar rhs) (EConstr.of_constr lhs)) then let id' = destVar rhs in subst_on l2r id' lhs, early_clear id' thin else @@ -2763,8 +2769,8 @@ let generalized_name c t ids cl = function let generalize_goal_gen env sigma ids i ((occs,c,b),na) t cl = let open Context.Rel.Declaration in let decls,cl = decompose_prod_n_assum i cl in - let dummy_prod = it_mkProd_or_LetIn mkProp decls in - let newdecls,_ = decompose_prod_n_assum i (subst_term_gen eq_constr_nounivs c dummy_prod) in + let dummy_prod = EConstr.of_constr (it_mkProd_or_LetIn mkProp decls) in + let newdecls,_ = decompose_prod_n_assum i (subst_term_gen sigma EConstr.eq_constr_nounivs (EConstr.of_constr c) dummy_prod) in let cl',sigma' = subst_closed_term_occ env sigma (AtOccs occs) c (it_mkProd_or_LetIn cl newdecls) in let na = generalized_name c t ids cl' na in let decl = match b with @@ -2782,10 +2788,11 @@ let generalize_goal gl i ((occs,c,b),na as o) (cl,sigma) = let old_generalize_dep ?(with_let=false) c gl = let env = pf_env gl in let sign = pf_hyps gl in + let sigma = project gl in let init_ids = ids_of_named_context (Global.named_context()) in let seek (d:Context.Named.Declaration.t) (toquant:Context.Named.t) = - if List.exists (fun d' -> occur_var_in_decl env (NamedDecl.get_id d') d) toquant - || dependent_in_decl c d then + if List.exists (fun d' -> occur_var_in_decl env sigma (NamedDecl.get_id d') d) toquant + || dependent_in_decl sigma (EConstr.of_constr c) d then d::toquant else toquant in @@ -2901,14 +2908,14 @@ let specialize (c,lbind) ipat = let (thd,tstack) = whd_nored_stack clause.evd (clenv_value clause) in let rec chk = function | [] -> [] - | t::l -> if occur_meta t then [] else t :: chk l + | t::l -> if occur_meta clause.evd (EConstr.of_constr t) then [] else t :: chk l in let tstack = chk tstack in let term = applist(thd,List.map (nf_evar clause.evd) tstack) in - if occur_meta term then + if occur_meta clause.evd (EConstr.of_constr term) then user_err (str "Cannot infer an instance for " ++ - pr_name (meta_name clause.evd (List.hd (collect_metas term))) ++ + pr_name (meta_name clause.evd (List.hd (collect_metas clause.evd (EConstr.of_constr term)))) ++ str "."); clause.evd, term in let typ = Retyping.get_type_of env sigma term in @@ -3143,10 +3150,11 @@ let induct_discharge with_evars dests avoid' tac (avoid,ra) names = let expand_projections env sigma c = let sigma = Sigma.to_evar_map sigma in let rec aux env c = - match kind_of_term c with - | Proj (p, c) -> Retyping.expand_projection env sigma p (aux env c) [] - | _ -> map_constr_with_full_binders push_rel aux env c - in aux env c + match EConstr.kind sigma c with + | Proj (p, c) -> EConstr.of_constr (Retyping.expand_projection env sigma p (EConstr.Unsafe.to_constr (aux env c)) []) + | _ -> map_constr_with_full_binders sigma push_rel aux env c + in + EConstr.Unsafe.to_constr (aux env (EConstr.of_constr c)) (* Marche pas... faut prendre en compte l'occurrence précise... *) @@ -3173,16 +3181,17 @@ let atomize_param_of_ind_then (indref,nparams,_) hyp0 tac = else let c = List.nth argl (i-1) in match kind_of_term c with - | Var id when not (List.exists (occur_var env id) args') && - not (List.exists (occur_var env id) params') -> + | Var id when not (List.exists (fun c -> occur_var env (Sigma.to_evar_map sigma) id (EConstr.of_constr c)) args') && + not (List.exists (fun c -> occur_var env (Sigma.to_evar_map sigma) id (EConstr.of_constr c)) params') -> (* Based on the knowledge given by the user, all constraints on the variable are generalizable in the current environment so that it is clearable after destruction *) atomize_one (i-1) (c::args) (c::args') (id::avoid) | _ -> let c' = expand_projections env' sigma c in - if List.exists (dependent c) params' || - List.exists (dependent c) args' + let dependent t = dependent (Sigma.to_evar_map sigma) (EConstr.of_constr c) (EConstr.of_constr t) in + if List.exists dependent params' || + List.exists dependent args' then (* This is a case where the argument is constrained in a way which would require some kind of inversion; we @@ -3272,7 +3281,7 @@ let atomize_param_of_ind_then (indref,nparams,_) hyp0 tac = exception Shunt of Id.t move_location -let cook_sign hyp0_opt inhyps indvars env = +let cook_sign hyp0_opt inhyps indvars env sigma = (* First phase from L to R: get [toclear], [decldep] and [statuslist] for the hypotheses before (= more ancient than) hyp0 (see above) *) let toclear = ref [] in @@ -3299,11 +3308,11 @@ let cook_sign hyp0_opt inhyps indvars env = rhyp end else let dephyp0 = List.is_empty inhyps && - (Option.cata (fun id -> occur_var_in_decl env id decl) false hyp0_opt) + (Option.cata (fun id -> occur_var_in_decl env sigma id decl) false hyp0_opt) in let depother = List.is_empty inhyps && - (List.exists (fun id -> occur_var_in_decl env id decl) indvars || - List.exists (fun decl' -> occur_var_in_decl env (NamedDecl.get_id decl') decl) !decldeps) + (List.exists (fun id -> occur_var_in_decl env sigma id decl) indvars || + List.exists (fun decl' -> occur_var_in_decl env sigma (NamedDecl.get_id decl') decl) !decldeps) in if not (List.is_empty inhyps) && Id.List.mem hyp inhyps || dephyp0 || depother @@ -3549,7 +3558,7 @@ let make_abstract_generalize env id typ concl dep ctx body c eqs args refls = Sigma (mkApp (appeqs, abshypt), sigma, p) end } -let hyps_of_vars env sign nogen hyps = +let hyps_of_vars env sigma sign nogen hyps = if Id.Set.is_empty hyps then [] else let (_,lh) = @@ -3559,7 +3568,7 @@ let hyps_of_vars env sign nogen hyps = if Id.Set.mem x nogen then (hs,hl) else if Id.Set.mem x hs then (hs,x::hl) else - let xvars = global_vars_set_of_decl env d in + let xvars = global_vars_set_of_decl env sigma d in if not (Id.Set.is_empty (Id.Set.diff xvars hs)) then (Id.Set.add x hs, x :: hl) else (hs, hl)) @@ -3592,7 +3601,7 @@ let abstract_args gl generalize_vars dep id defined f args = let sigma = ref (Tacmach.project gl) in let env = Tacmach.pf_env gl in let concl = Tacmach.pf_concl gl in - let dep = dep || dependent (mkVar id) concl in + let dep = dep || local_occur_var !sigma id (EConstr.of_constr concl) in let avoid = ref [] in let get_id name = let id = fresh_id !avoid (match name with Name n -> n | Anonymous -> Id.of_string "gen_x") gl in @@ -3659,7 +3668,7 @@ let abstract_args gl generalize_vars dep id defined f args = let vars = if generalize_vars then let nogen = Id.Set.add id nogen in - hyps_of_vars (pf_env gl) (pf_hyps gl) nogen vars + hyps_of_vars (pf_env gl) (project gl) (pf_hyps gl) nogen vars else [] in let body, c' = @@ -3845,7 +3854,7 @@ let compute_elim_sig ?elimc elimt = let ccl = exchange_hd_app (mkVar (Id.of_string "__QI_DUMMY__")) conclusion in let concl_with_args = it_mkProd_or_LetIn ccl args_indargs in - let nparams = Int.Set.cardinal (free_rels concl_with_args) in + let nparams = Int.Set.cardinal (free_rels Evd.empty (** FIXME *) (EConstr.of_constr concl_with_args)) in let preds,params = List.chop (List.length params_preds - nparams) params_preds in (* A first approximation, further analysis will tweak it *) @@ -3905,7 +3914,7 @@ let compute_elim_sig ?elimc elimt = with e when CErrors.noncritical e -> error "Cannot find the inductive type of the inductive scheme." -let compute_scheme_signature scheme names_info ind_type_guess = +let compute_scheme_signature evd scheme names_info ind_type_guess = let open Context.Rel.Declaration in let f,l = decompose_app scheme.concl in (* Vérifier que les arguments de Qi sont bien les xi. *) @@ -3940,9 +3949,9 @@ let compute_scheme_signature scheme names_info ind_type_guess = let rec check_branch p c = match kind_of_term c with | Prod (_,t,c) -> - (is_pred p t, true, dependent (mkRel 1) c) :: check_branch (p+1) c + (is_pred p t, true, not (EConstr.Vars.noccurn evd 1 (EConstr.of_constr c))) :: check_branch (p+1) c | LetIn (_,_,_,c) -> - (OtherArg, false, dependent (mkRel 1) c) :: check_branch (p+1) c + (OtherArg, false, not (EConstr.Vars.noccurn evd 1 (EConstr.of_constr c))) :: check_branch (p+1) c | _ when is_pred p c == IndArg -> [] | _ -> raise Exit in @@ -3975,7 +3984,7 @@ let compute_scheme_signature scheme names_info ind_type_guess = different. *) let compute_elim_signature (evd,(elimc,elimt),ind_type_guess) names_info = let scheme = compute_elim_sig ~elimc:elimc elimt in - evd, (compute_scheme_signature scheme names_info ind_type_guess, scheme) + evd, (compute_scheme_signature evd scheme names_info ind_type_guess, scheme) let guess_elim isrec dep s hyp0 gl = let tmptyp0 = Tacmach.New.pf_get_hyp_typ hyp0 gl in @@ -4022,7 +4031,7 @@ let find_induction_type isrec elim hyp0 gl = let evd, (elimc,elimt),ind_guess = given_elim hyp0 e gl in let scheme = compute_elim_sig ~elimc elimt in if Option.is_empty scheme.indarg then error "Cannot find induction type"; - let indsign = compute_scheme_signature scheme hyp0 ind_guess in + let indsign = compute_scheme_signature evd scheme hyp0 ind_guess in let elim = ({elimindex = Some(-1); elimbody = elimc; elimrename = None},elimt) in scheme, ElimUsing (elim,indsign) in @@ -4049,7 +4058,7 @@ let get_eliminator elim dep s gl = | ElimOver (isrec,id) -> let evd, (elimc,elimt),_ as elims = guess_elim isrec dep s id gl in let _, (l, s) = compute_elim_signature elims id in - let branchlengthes = List.map (fun d -> assert (RelDecl.is_local_assum d); pi1 (decompose_prod_letin (RelDecl.get_type d))) + let branchlengthes = List.map (fun d -> assert (RelDecl.is_local_assum d); pi1 (decompose_prod_letin (Tacmach.New.project gl) (EConstr.of_constr (RelDecl.get_type d)))) (List.rev s.branches) in evd, isrec, ({elimindex = None; elimbody = elimc; elimrename = Some (isrec,Array.of_list branchlengthes)}, elimt), l @@ -4118,8 +4127,8 @@ let apply_induction_in_context with_evars hyp0 inhyps elim indvars names induct_ let env = Proofview.Goal.env gl in let sigma = Sigma.to_evar_map sigma in let concl = Tacmach.New.pf_nf_concl gl in - let statuslists,lhyp0,toclear,deps,avoid,dep_in_hyps = cook_sign hyp0 inhyps indvars env in - let dep_in_concl = Option.cata (fun id -> occur_var env id concl) false hyp0 in + let statuslists,lhyp0,toclear,deps,avoid,dep_in_hyps = cook_sign hyp0 inhyps indvars env sigma in + let dep_in_concl = Option.cata (fun id -> occur_var env sigma id (EConstr.of_constr concl)) false hyp0 in let dep = dep_in_hyps || dep_in_concl in let tmpcl = it_mkNamedProd_or_LetIn concl deps in let s = Retyping.get_sort_family_of env sigma tmpcl in @@ -4207,7 +4216,7 @@ let induction_without_atomization isrec with_evars elim names lid = (* assume that no occurrences are selected *) let clear_unselected_context id inhyps cls = Proofview.Goal.nf_enter { enter = begin fun gl -> - if occur_var (Tacmach.New.pf_env gl) id (Tacmach.New.pf_concl gl) && + if occur_var (Tacmach.New.pf_env gl) (Tacmach.New.project gl) id (EConstr.of_constr (Tacmach.New.pf_concl gl)) && cls.concl_occs == NoOccurrences then user_err (str "Conclusion must be mentioned: it depends on " ++ pr_id id @@ -4219,7 +4228,7 @@ let clear_unselected_context id inhyps cls = if Id.List.mem id' inhyps then (* if selected, do not erase *) None else (* erase if not selected and dependent on id or selected hyps *) - let test id = occur_var_in_decl (Tacmach.New.pf_env gl) id d in + let test id = occur_var_in_decl (Tacmach.New.pf_env gl) (Tacmach.New.project gl) id d in if List.exists test (id::inhyps) then Some id' else None in let ids = List.map_filter to_erase (Proofview.Goal.hyps gl) in clear ids @@ -4246,7 +4255,7 @@ let use_bindings env sigma elim must_be_closed (c,lbind) typ = let rec find_clause typ = try let indclause = make_clenv_binding env sigma (c,typ) lbind in - if must_be_closed && occur_meta (clenv_value indclause) then + if must_be_closed && occur_meta indclause.evd (EConstr.of_constr (clenv_value indclause)) then error "Need a fully applied argument."; (* We lose the possibility of coercions in with-bindings *) let (sigma, c) = pose_all_metas_as_evars env indclause.evd (clenv_value indclause) in @@ -4351,10 +4360,10 @@ let pose_induction_arg_then isrec with_evars (is_arg_pure_hyp,from_prefix) elim Sigma (tac, sigma', p +> q) end } -let has_generic_occurrences_but_goal cls id env ccl = +let has_generic_occurrences_but_goal cls id env sigma ccl = clause_with_generic_context_selection cls && (* TODO: whd_evar of goal *) - (cls.concl_occs != NoOccurrences || not (occur_var env id ccl)) + (cls.concl_occs != NoOccurrences || not (occur_var env sigma id (EConstr.of_constr ccl))) let induction_gen clear_flag isrec with_evars elim ((_pending,(c,lbind)),(eqname,names) as arg) cls = @@ -4371,7 +4380,7 @@ let induction_gen clear_flag isrec with_evars elim isVar c && not (mem_named_context_val (destVar c) (Global.named_context_val ())) && lbind == NoBindings && not with_evars && Option.is_empty eqname && clear_flag == None - && has_generic_occurrences_but_goal cls (destVar c) env ccl in + && has_generic_occurrences_but_goal cls (destVar c) env (Sigma.to_evar_map sigma) ccl in let enough_applied = check_enough_applied env sigma elim t in if is_arg_pure_hyp && enough_applied then (* First case: induction on a variable already in an inductive type and @@ -4423,11 +4432,12 @@ let induction_gen_l isrec with_evars elim names lc = | _ -> Proofview.Goal.enter { enter = begin fun gl -> let type_of = Tacmach.New.pf_unsafe_type_of gl in + let sigma = Tacmach.New.project gl in let x = id_of_name_using_hdchar (Global.env()) (type_of c) Anonymous in let id = new_fresh_id [] x gl in - let newl' = List.map (replace_term c (mkVar id)) l' in + let newl' = List.map (fun r -> replace_term sigma (EConstr.of_constr c) (EConstr.mkVar id) (EConstr.of_constr r)) l' in let _ = newlc:=id::!newlc in Tacticals.New.tclTHEN (letin_tac None (Name id) c None allHypsAndConcl) @@ -4601,8 +4611,9 @@ let reflexivity_red allowred = (* PL: usual reflexivity don't perform any reduction when searching for an equality, but we may need to do some when called back from inside setoid_reflexivity (see Optimize cases in setoid_replace.ml). *) + let sigma = Tacmach.New.project gl in let concl = maybe_betadeltaiota_concl allowred gl in - match match_with_equality_type concl with + match match_with_equality_type sigma concl with | None -> Proofview.tclZERO NoEquationFound | Some _ -> one_constructor 1 NoBindings end } diff --git a/tactics/term_dnet.ml b/tactics/term_dnet.ml index e4b45489d..6294f9fdc 100644 --- a/tactics/term_dnet.ml +++ b/tactics/term_dnet.ml @@ -351,7 +351,7 @@ struct (fun id acc -> let c_id = Opt.reduce (Ident.constr_of id) in let (ctx,wc) = - try Termops.align_prod_letin whole_c c_id + try Termops.align_prod_letin Evd.empty (EConstr.of_constr whole_c) (EConstr.of_constr c_id) (** FIXME *) with Invalid_argument _ -> [],c_id in let wc,whole_c = if Opt.direction then whole_c,wc else wc,whole_c in try diff --git a/toplevel/assumptions.ml b/toplevel/assumptions.ml index 8865cd646..deb2ed3e0 100644 --- a/toplevel/assumptions.ml +++ b/toplevel/assumptions.ml @@ -144,6 +144,27 @@ let label_of = function | ConstructRef ((kn,_),_) -> pi3 (repr_mind kn) | VarRef id -> Label.of_id id +let fold_constr_with_full_binders g f n acc c = + let open Context.Rel.Declaration in + match kind_of_term 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 + | 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 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 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 rec traverse current ctx accu t = match kind_of_term t with | Var id -> let body () = id |> Global.lookup_named |> NamedDecl.get_value in @@ -166,10 +187,10 @@ let rec traverse current ctx accu t = match kind_of_term t with traverse_object ~inhabits:(current,ctx,Vars.subst1 mkProp oty) accu body (ConstRef kn) | _ -> - Termops.fold_constr_with_full_binders + fold_constr_with_full_binders Context.Rel.add (traverse current) ctx accu t end -| _ -> Termops.fold_constr_with_full_binders +| _ -> fold_constr_with_full_binders Context.Rel.add (traverse current) ctx accu t and traverse_object ?inhabits (curr, data, ax2ty) body obj = diff --git a/toplevel/class.ml b/toplevel/class.ml index 0dc799014..46854e5c0 100644 --- a/toplevel/class.ml +++ b/toplevel/class.ml @@ -85,7 +85,7 @@ let uniform_cond nargs lt = let rec aux = function | (0,[]) -> true | (n,t::l) -> - let t = strip_outer_cast t in + let t = strip_outer_cast Evd.empty (EConstr.of_constr t) (** FIXME *) in isRel t && Int.equal (destRel t) n && aux ((n-1),l) | _ -> false in diff --git a/toplevel/command.ml b/toplevel/command.ml index ef918ef8d..62bbd4b97 100644 --- a/toplevel/command.ml +++ b/toplevel/command.ml @@ -685,7 +685,7 @@ let is_recursive mie = let rec is_recursive_constructor lift typ = match Term.kind_of_term typ with | Prod (_,arg,rest) -> - Termops.dependent (mkRel lift) arg || + not (EConstr.Vars.noccurn Evd.empty (** FIXME *) lift (EConstr.of_constr arg)) || is_recursive_constructor (lift+1) rest | LetIn (na,b,t,rest) -> is_recursive_constructor (lift+1) rest | _ -> false @@ -813,11 +813,11 @@ let warn_non_full_mutual = (fun (x,xge,y,yge,isfix,rest) -> non_full_mutual_message x xge y yge isfix rest) -let check_mutuality env isfix fixl = +let check_mutuality env evd isfix fixl = let names = List.map fst fixl in let preorder = List.map (fun (id,def) -> - (id, List.filter (fun id' -> not (Id.equal id id') && occur_var env id' def) names)) + (id, List.filter (fun id' -> not (Id.equal id id') && occur_var env evd id' (EConstr.of_constr def)) names)) fixl in let po = partial_order Id.equal preorder in match List.filter (function (_,Inr _) -> true | _ -> false) po with @@ -1144,7 +1144,7 @@ let check_recursive isfix env evd (fixnames,fixdefs,_) = check_evars_are_solved env evd (Evd.empty,evd); if List.for_all Option.has_some fixdefs then begin let fixdefs = List.map Option.get fixdefs in - check_mutuality env isfix (List.combine fixnames fixdefs) + check_mutuality env evd isfix (List.combine fixnames fixdefs) end let interp_fixpoint l ntns = diff --git a/toplevel/command.mli b/toplevel/command.mli index 616afb91f..fae783ef0 100644 --- a/toplevel/command.mli +++ b/toplevel/command.mli @@ -170,7 +170,7 @@ val do_cofixpoint : (** Utils *) -val check_mutuality : Environ.env -> bool -> (Id.t * types) list -> unit +val check_mutuality : Environ.env -> Evd.evar_map -> bool -> (Id.t * types) list -> unit val declare_fix : ?opaque:bool -> definition_kind -> Universes.universe_binders -> Univ.universe_context -> Id.t -> Safe_typing.private_constants Entries.proof_output -> types -> Impargs.manual_implicits -> global_reference diff --git a/toplevel/himsg.ml b/toplevel/himsg.ml index 891662b93..bfe86053e 100644 --- a/toplevel/himsg.ml +++ b/toplevel/himsg.ml @@ -1090,7 +1090,7 @@ let error_ill_formed_inductive env c v = let error_ill_formed_constructor env id c v nparams nargs = let pv = pr_lconstr_env env Evd.empty v in - let atomic = Int.equal (nb_prod c) 0 in + let atomic = Int.equal (nb_prod Evd.empty (EConstr.of_constr c)) (** FIXME *) 0 in str "The type of constructor" ++ brk(1,1) ++ pr_id id ++ brk(1,1) ++ str "is not valid;" ++ brk(1,1) ++ strbrk (if atomic then "it must be " else "its conclusion must be ") ++ diff --git a/toplevel/indschemes.ml b/toplevel/indschemes.ml index 48521a8e5..3f818f960 100644 --- a/toplevel/indschemes.ml +++ b/toplevel/indschemes.ml @@ -313,7 +313,7 @@ let warn_cannot_build_congruence = strbrk "Cannot build congruence scheme because eq is not found") let declare_congr_scheme ind = - if Hipattern.is_equality_type (mkInd ind) then begin + if Hipattern.is_equality_type Evd.empty (mkInd ind) (** FIXME *) then begin if try Coqlib.check_required_library Coqlib.logic_module_name; true with e when CErrors.noncritical e -> false @@ -472,7 +472,7 @@ let build_combined_scheme env schemes = let (c, t) = List.hd defs in let ctx, ind, nargs = find_inductive t in (* Number of clauses, including the predicates quantification *) - let prods = nb_prod t - (nargs + 1) in + let prods = nb_prod Evd.empty (EConstr.of_constr t) - (nargs + 1) (** FIXME *) in let coqand = Coqlib.build_coq_and () and coqconj = Coqlib.build_coq_conj () in let relargs = rel_vect 0 prods in let concls = List.rev_map diff --git a/toplevel/obligations.ml b/toplevel/obligations.ml index 9ada04317..1d5e4a2fa 100644 --- a/toplevel/obligations.ml +++ b/toplevel/obligations.ml @@ -145,7 +145,7 @@ let rec chop_product n t = if Int.equal n 0 then Some t else match kind_of_term t with - | Prod (_, _, b) -> if noccurn 1 b then chop_product (pred n) (Termops.pop b) else None + | Prod (_, _, b) -> if noccurn 1 b then chop_product (pred n) (Termops.pop (EConstr.of_constr b)) else None | _ -> None let evar_dependencies evm oev = @@ -396,7 +396,7 @@ let subst_deps expand obls deps t = (Vars.replace_vars (List.map (fun (n, (_, b)) -> n, b) osubst) t) let rec prod_app t n = - match kind_of_term (Termops.strip_outer_cast t) with + match kind_of_term (Termops.strip_outer_cast Evd.empty (EConstr.of_constr t)) (** FIXME *) with | Prod (_,_,b) -> subst1 n b | LetIn (_, b, t, b') -> prod_app (subst1 b b') n | _ -> @@ -521,7 +521,7 @@ let compute_possible_guardness_evidences (n,_) fixbody fixtype = but doing it properly involves delta-reduction, and it finally doesn't seem to worth the effort (except for huge mutual fixpoints ?) *) - let m = Termops.nb_prod fixtype in + let m = Termops.nb_prod Evd.empty (EConstr.of_constr fixtype) (** FIXME *) in let ctx = fst (decompose_prod_n_assum m fixtype) in List.map_i (fun i _ -> i) 0 ctx diff --git a/toplevel/record.ml b/toplevel/record.ml index c43b32762..7dee4aae0 100644 --- a/toplevel/record.ml +++ b/toplevel/record.ml @@ -374,7 +374,7 @@ let structure_signature ctx = let evm = Sigma.to_evar_map evm in let new_tl = Util.List.map_i (fun pos decl -> - RelDecl.map_type (fun t -> Termops.replace_term (mkRel pos) (mkEvar(ev,[||])) t) decl) 1 tl in + RelDecl.map_type (fun t -> Termops.replace_term evm (EConstr.mkRel pos) (EConstr.mkEvar(ev,[||])) (EConstr.of_constr t)) decl) 1 tl in deps_to_evar evm new_tl in deps_to_evar Evd.empty (List.rev ctx) diff --git a/toplevel/search.ml b/toplevel/search.ml index d319b2419..c0bcc403c 100644 --- a/toplevel/search.ml +++ b/toplevel/search.ml @@ -112,7 +112,7 @@ let generic_search glnumopt fn = (** This function tries to see whether the conclusion matches a pattern. *) (** FIXME: this is quite dummy, we may find a more efficient algorithm. *) let rec pattern_filter pat ref env typ = - let typ = Termops.strip_outer_cast typ in + let typ = Termops.strip_outer_cast Evd.empty (EConstr.of_constr typ) (** FIXME *) in if Constr_matching.is_matching env Evd.empty pat typ then true else match kind_of_term typ with | Prod (_, _, typ) @@ -120,7 +120,7 @@ let rec pattern_filter pat ref env typ = | _ -> false let rec head_filter pat ref env typ = - let typ = Termops.strip_outer_cast typ in + let typ = Termops.strip_outer_cast Evd.empty (EConstr.of_constr typ) (** FIXME *) in if Constr_matching.is_matching_head env Evd.empty pat typ then true else match kind_of_term typ with | Prod (_, _, typ) diff --git a/toplevel/vernacentries.ml b/toplevel/vernacentries.ml index ede88399e..bbbdbdb67 100644 --- a/toplevel/vernacentries.ml +++ b/toplevel/vernacentries.ml @@ -107,7 +107,7 @@ let show_intro all = let {Evd.it=gls ; sigma=sigma; } = Proof.V82.subgoals pf in if not (List.is_empty gls) then begin let gl = {Evd.it=List.hd gls ; sigma = sigma; } in - let l,_= decompose_prod_assum (Termops.strip_outer_cast (pf_concl gl)) in + let l,_= decompose_prod_assum (Termops.strip_outer_cast (project gl) (EConstr.of_constr (pf_concl gl))) in if all then let lid = Tactics.find_intro_names l gl in Feedback.msg_notice (hov 0 (prlist_with_sep spc pr_id lid)) -- cgit v1.2.3 From 8f6aab1f4d6d60842422abc5217daac806eb0897 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 1 Nov 2016 20:53:32 +0100 Subject: Reductionops API using EConstr. --- engine/termops.mli | 2 +- interp/notation.ml | 2 +- ltac/extratactics.ml4 | 4 +- ltac/rewrite.ml | 47 +-- ltac/tacinterp.ml | 2 +- ltac/tactic_matching.ml | 2 +- plugins/decl_mode/decl_proof_instr.ml | 4 +- plugins/extraction/extraction.ml | 22 +- plugins/firstorder/instances.ml | 2 +- plugins/firstorder/unify.ml | 4 +- plugins/funind/functional_principles_proofs.ml | 24 +- plugins/funind/functional_principles_types.ml | 2 +- plugins/funind/glob_term_to_relation.ml | 8 +- plugins/funind/invfun.ml | 14 +- plugins/funind/merge.ml | 2 +- plugins/funind/recdef.ml | 4 +- plugins/micromega/coq_micromega.ml | 2 +- plugins/omega/coq_omega.ml | 2 +- plugins/quote/quote.ml | 4 +- plugins/setoid_ring/newring.ml | 5 +- plugins/ssrmatching/ssrmatching.ml4 | 4 +- pretyping/cases.ml | 24 +- pretyping/classops.ml | 23 +- pretyping/coercion.ml | 24 +- pretyping/constr_matching.ml | 2 +- pretyping/detyping.ml | 2 +- pretyping/evarconv.ml | 177 +++++----- pretyping/evarconv.mli | 8 +- pretyping/evardefine.ml | 10 +- pretyping/evarsolve.ml | 16 +- pretyping/evarsolve.mli | 2 +- pretyping/indrec.ml | 16 +- pretyping/inductiveops.ml | 53 +-- pretyping/inductiveops.mli | 18 +- pretyping/nativenorm.ml | 95 +++--- pretyping/pretyping.ml | 20 +- pretyping/recordops.ml | 5 +- pretyping/recordops.mli | 2 +- pretyping/reductionops.ml | 445 ++++++++++++++----------- pretyping/reductionops.mli | 100 +++--- pretyping/retyping.ml | 29 +- pretyping/tacred.ml | 266 ++++++++------- pretyping/typeclasses.ml | 2 +- pretyping/typing.ml | 12 +- pretyping/unification.ml | 66 ++-- pretyping/vnorm.ml | 110 +++--- pretyping/vnorm.mli | 2 +- printing/prettyp.ml | 4 +- proofs/clenv.ml | 10 +- proofs/logic.ml | 16 +- proofs/redexpr.ml | 7 +- proofs/tacmach.ml | 24 +- proofs/tacmach.mli | 6 +- stm/lemmas.ml | 6 +- tactics/autorewrite.ml | 2 +- tactics/contradiction.ml | 10 +- tactics/eauto.ml | 2 +- tactics/elim.ml | 4 +- tactics/eqdecide.ml | 2 +- tactics/eqschemes.ml | 4 +- tactics/equality.ml | 60 ++-- tactics/hints.ml | 4 +- tactics/hipattern.ml | 4 +- tactics/inv.ml | 2 +- tactics/leminv.ml | 6 +- tactics/tactics.ml | 81 ++--- tactics/tactics.mli | 2 +- toplevel/auto_ind_decl.ml | 10 +- toplevel/class.ml | 4 +- toplevel/command.ml | 8 +- toplevel/himsg.ml | 6 +- toplevel/obligations.ml | 6 +- toplevel/record.ml | 2 +- toplevel/vernacentries.ml | 2 +- 74 files changed, 1050 insertions(+), 935 deletions(-) diff --git a/engine/termops.mli b/engine/termops.mli index 5d53ce09e..b536b0fb8 100644 --- a/engine/termops.mli +++ b/engine/termops.mli @@ -17,7 +17,7 @@ open Environ (** printers *) val print_sort : sorts -> std_ppcmds val pr_sort_family : sorts_family -> std_ppcmds -val pr_fix : (constr -> std_ppcmds) -> fixpoint -> std_ppcmds +val pr_fix : ('a -> std_ppcmds) -> ('a, 'a) pfixpoint -> std_ppcmds (** debug printer: do not use to display terms to the casual user... *) val set_print_constr : (env -> constr -> std_ppcmds) -> unit diff --git a/interp/notation.ml b/interp/notation.ml index 948d624a2..88ae4695b 100644 --- a/interp/notation.ml +++ b/interp/notation.ml @@ -615,7 +615,7 @@ let find_scope_class_opt = function (* Special scopes associated to arguments of a global reference *) let rec compute_arguments_classes t = - match kind_of_term (Reductionops.whd_betaiotazeta Evd.empty t) with + match kind_of_term (Reductionops.whd_betaiotazeta Evd.empty (EConstr.of_constr t)) with | Prod (_,t,u) -> let cl = try Some (compute_scope_class t) with Not_found -> None in cl :: compute_arguments_classes u diff --git a/ltac/extratactics.ml4 b/ltac/extratactics.ml4 index beaf778a6..2405e3c42 100644 --- a/ltac/extratactics.ml4 +++ b/ltac/extratactics.ml4 @@ -307,7 +307,7 @@ let project_hint pri l2r r = | _ -> assert false in let p = if l2r then build_coq_iff_left_proj () else build_coq_iff_right_proj () in - let c = Reductionops.whd_beta Evd.empty (mkApp (c, Context.Rel.to_extended_vect 0 sign)) in + let c = Reductionops.whd_beta Evd.empty (EConstr.of_constr (mkApp (c, Context.Rel.to_extended_vect 0 sign))) in let c = it_mkLambda_or_LetIn (mkApp (p,[|mkArrow a (lift 1 b);mkArrow b (lift 1 a);c|])) sign in let id = @@ -738,7 +738,7 @@ let mkCaseEq a : unit Proofview.tactic = let concl = Proofview.Goal.concl gl in let env = Proofview.Goal.env gl in (** FIXME: this looks really wrong. Does anybody really use this tactic? *) - let Sigma (c, _, _) = (Tacred.pattern_occs [Locus.OnlyOccurrences [1], a]).Reductionops.e_redfun env (Sigma.Unsafe.of_evar_map Evd.empty) concl in + let Sigma (c, _, _) = (Tacred.pattern_occs [Locus.OnlyOccurrences [1], a]).Reductionops.e_redfun env (Sigma.Unsafe.of_evar_map Evd.empty) (EConstr.of_constr concl) in change_concl c end }; simplest_case a] diff --git a/ltac/rewrite.ml b/ltac/rewrite.ml index cd76d4746..7ef3ace53 100644 --- a/ltac/rewrite.ml +++ b/ltac/rewrite.ml @@ -221,12 +221,12 @@ end) = struct | Some (x, Some rel) -> evars, rel in let rec aux env evars ty l = - let t = Reductionops.whd_all env (goalevars evars) ty in + let t = Reductionops.whd_all env (goalevars evars) (EConstr.of_constr ty) in match kind_of_term t, l with | Prod (na, ty, b), obj :: cstrs -> - let b = Reductionops.nf_betaiota (goalevars evars) b in + let b = Reductionops.nf_betaiota (goalevars evars) (EConstr.of_constr b) in if noccurn 1 b (* non-dependent product *) then - let ty = Reductionops.nf_betaiota (goalevars evars) ty in + let ty = Reductionops.nf_betaiota (goalevars evars) (EConstr.of_constr ty) in let (evars, b', arg, cstrs) = aux env evars (subst1 mkProp b) cstrs in let evars, relty = mk_relty evars env ty obj in let evars, newarg = app_poly env evars respectful [| ty ; b' ; relty ; arg |] in @@ -235,7 +235,7 @@ end) = struct let (evars, b, arg, cstrs) = aux (Environ.push_rel (LocalAssum (na, ty)) env) evars b cstrs in - let ty = Reductionops.nf_betaiota (goalevars evars) ty in + let ty = Reductionops.nf_betaiota (goalevars evars) (EConstr.of_constr ty) in let pred = mkLambda (na, ty, b) in let liftarg = mkLambda (na, ty, arg) in let evars, arg' = app_poly env evars forall_relation [| ty ; pred ; liftarg |] in @@ -245,7 +245,7 @@ end) = struct | _, [] -> (match finalcstr with | None | Some (_, None) -> - let t = Reductionops.nf_betaiota (fst evars) ty in + let t = Reductionops.nf_betaiota (fst evars) (EConstr.of_constr ty) in let evars, rel = mk_relty evars env t None in evars, t, rel, [t, Some rel] | Some (t, Some rel) -> evars, t, rel, [t, Some rel]) @@ -286,23 +286,24 @@ end) = struct else (* None in Prop, use arrow *) (app_poly env evd arrow [| a; b |]), unfold_impl - let rec decomp_pointwise n c = + let rec decomp_pointwise sigma n c = if Int.equal n 0 then c else + let open EConstr in match kind_of_term c with | App (f, [| a; b; relb |]) when Globnames.is_global (pointwise_relation_ref ()) f -> - decomp_pointwise (pred n) relb + decomp_pointwise sigma (pred n) relb | App (f, [| a; b; arelb |]) when Globnames.is_global (forall_relation_ref ()) f -> - decomp_pointwise (pred n) (Reductionops.beta_applist (arelb, [mkRel 1])) + decomp_pointwise sigma (pred n) (Reductionops.beta_applist sigma (EConstr.of_constr arelb, [mkRel 1])) | _ -> invalid_arg "decomp_pointwise" - let rec apply_pointwise rel = function + let rec apply_pointwise sigma rel = function | arg :: args -> (match kind_of_term rel with | App (f, [| a; b; relb |]) when Globnames.is_global (pointwise_relation_ref ()) f -> - apply_pointwise relb args + apply_pointwise sigma relb args | App (f, [| a; b; arelb |]) when Globnames.is_global (forall_relation_ref ()) f -> - apply_pointwise (Reductionops.beta_applist (arelb, [arg])) args + apply_pointwise sigma (Reductionops.beta_applist sigma (EConstr.of_constr arelb, [EConstr.of_constr arg])) args | _ -> invalid_arg "apply_pointwise") | [] -> rel @@ -348,7 +349,7 @@ end) = struct let unlift_cstr env sigma = function | None -> None - | Some codom -> Some (decomp_pointwise 1 codom) + | Some codom -> Some (decomp_pointwise (goalevars sigma) 1 codom) (** Looking up declared rewrite relations (instances of [RewriteRelation]) *) let is_applied_rewrite_relation env sigma rels t = @@ -430,7 +431,7 @@ module TypeGlobal = struct end let sort_of_rel env evm rel = - Reductionops.sort_of_arity env evm (Retyping.get_type_of env evm rel) + Reductionops.sort_of_arity env evm (EConstr.of_constr (Retyping.get_type_of env evm rel)) let is_applied_rewrite_relation = PropGlobal.is_applied_rewrite_relation @@ -460,7 +461,7 @@ let evd_convertible env evd x y = with e when CErrors.noncritical e -> None let convertible env evd x y = - Reductionops.is_conv_leq env evd x y + Reductionops.is_conv_leq env evd (EConstr.of_constr x) (EConstr.of_constr y) type hypinfo = { prf : constr; @@ -479,7 +480,7 @@ let error_no_relation () = error "Cannot find a relation to rewrite." let rec decompose_app_rel env evd t = (** Head normalize for compatibility with the old meta mechanism *) - let t = Reductionops.whd_betaiota evd t in + let t = Reductionops.whd_betaiota evd (EConstr.of_constr t) in match kind_of_term t with | App (f, [||]) -> assert false | App (f, [|arg|]) -> @@ -525,7 +526,7 @@ let decompose_applied_relation env sigma (c,l) = match find_rel ctype with | Some c -> c | None -> - let ctx,t' = Reductionops.splay_prod env sigma ctype in (* Search for underlying eq *) + let ctx,t' = Reductionops.splay_prod env sigma (EConstr.of_constr ctype) in (* Search for underlying eq *) match find_rel (it_mkProd_or_LetIn t' (List.map (fun (n,t) -> LocalAssum (n, t)) ctx)) with | Some c -> c | None -> error "Cannot find an homogeneous relation to rewrite." @@ -966,7 +967,7 @@ let unfold_match env sigma sk app = match kind_of_term app with | App (f', args) when eq_constant (fst (destConst f')) sk -> let v = Environ.constant_value_in (Global.env ()) (sk,Univ.Instance.empty)(*FIXME*) in - Reductionops.whd_beta sigma (mkApp (v, args)) + Reductionops.whd_beta sigma (EConstr.of_constr (mkApp (v, args))) | _ -> app let is_rew_cast = function RewCast _ -> true | _ -> false @@ -1055,11 +1056,11 @@ let subterm all flags (s : 'a pure_strategy) : 'a pure_strategy = let app = if prop then PropGlobal.apply_pointwise else TypeGlobal.apply_pointwise in - RewPrf (app rel argsl, mkApp (prf, args)) + RewPrf (app (goalevars evars) rel argsl, mkApp (prf, args)) | x -> x in let res = - { rew_car = Reductionops.hnf_prod_appvect env (goalevars evars) r.rew_car args; + { rew_car = Reductionops.hnf_prod_appvect env (goalevars evars) (EConstr.of_constr r.rew_car) (Array.map EConstr.of_constr args); rew_from = mkApp(r.rew_from, args); rew_to = mkApp(r.rew_to, args); rew_prf = prf; rew_evars = r.rew_evars } in @@ -1402,7 +1403,7 @@ module Strategies = fun { state = state ; env = env ; term1 = t ; ty1 = ty ; cstr = cstr ; evars = evars } -> let rfn, ckind = Redexpr.reduction_of_red_expr env r in let sigma = Sigma.Unsafe.of_evar_map (goalevars evars) in - let Sigma (t', sigma, _) = rfn.Reductionops.e_redfun env sigma t in + let Sigma (t', sigma, _) = rfn.Reductionops.e_redfun env sigma (EConstr.of_constr t) in let evars' = Sigma.to_evar_map sigma in if eq_constr t' t then state, Identity @@ -1417,7 +1418,7 @@ module Strategies = (* let sigma, (c,_) = Tacinterp.interp_open_constr_with_bindings is env (goalevars evars) c in *) let sigma, c = Pretyping.understand_tcc env (goalevars evars) c in let unfolded = - try Tacred.try_red_product env sigma c + try Tacred.try_red_product env sigma (EConstr.of_constr c) with e when CErrors.noncritical e -> error "fold: the term is not unfoldable !" in @@ -1511,7 +1512,7 @@ let cl_rewrite_clause_aux ?(abs=None) strat env avoid sigma concl is_hyp : resul let res = match res.rew_prf with | RewCast c -> None | RewPrf (rel, p) -> - let p = nf_zeta env evars' (Evarutil.nf_evar evars' p) in + let p = nf_zeta env evars' (EConstr.of_constr (Evarutil.nf_evar evars' p)) in let term = match abs with | None -> p @@ -1887,7 +1888,7 @@ let declare_projection n instance_id r = | _ -> typ in aux init in - let ctx,ccl = Reductionops.splay_prod_n (Global.env()) Evd.empty (3 * n) typ + let ctx,ccl = Reductionops.splay_prod_n env sigma (3 * n) (EConstr.of_constr typ) in it_mkProd_or_LetIn ccl ctx in let typ = it_mkProd_or_LetIn typ ctx in diff --git a/ltac/tacinterp.ml b/ltac/tacinterp.ml index e927ea064..34faa028f 100644 --- a/ltac/tacinterp.ml +++ b/ltac/tacinterp.ml @@ -783,7 +783,7 @@ let interp_may_eval f ist env sigma = function let (sigma,c_interp) = f ist env sigma c in let (redfun, _) = Redexpr.reduction_of_red_expr env redexp in let sigma = Sigma.Unsafe.of_evar_map sigma in - let Sigma (c, sigma, _) = redfun.Reductionops.e_redfun env sigma c_interp in + let Sigma (c, sigma, _) = redfun.Reductionops.e_redfun env sigma (EConstr.of_constr c_interp) in (Sigma.to_evar_map sigma, c) | ConstrContext ((loc,s),c) -> (try diff --git a/ltac/tactic_matching.ml b/ltac/tactic_matching.ml index ef45ee47e..43a58f6b2 100644 --- a/ltac/tactic_matching.ml +++ b/ltac/tactic_matching.ml @@ -84,7 +84,7 @@ let equal_instances env sigma (ctx',c') (ctx,c) = (* How to compare instances? Do we want the terms to be convertible? unifiable? Do we want the universe levels to be relevant? (historically, conv_x is used) *) - CList.equal Id.equal ctx ctx' && Reductionops.is_conv env sigma c' c + CList.equal Id.equal ctx ctx' && Reductionops.is_conv env sigma (EConstr.of_constr c') (EConstr.of_constr c) (** Merges two substitutions. Raises [Not_coherent_metas] when diff --git a/plugins/decl_mode/decl_proof_instr.ml b/plugins/decl_mode/decl_proof_instr.ml index 46fa5b408..c17c8dbb8 100644 --- a/plugins/decl_mode/decl_proof_instr.ml +++ b/plugins/decl_mode/decl_proof_instr.ml @@ -1092,7 +1092,7 @@ let thesis_for obj typ per_info env= ((Printer.pr_constr_env env Evd.empty obj) ++ spc () ++ str "cannot give an induction hypothesis (wrong parameters).") in let hd2 = (applist ((lift (List.length rc) per_info.per_pred),args@[obj])) in - compose_prod rc (Reductionops.whd_beta Evd.empty hd2) + compose_prod rc (Reductionops.whd_beta Evd.empty (EConstr.of_constr hd2)) let rec build_product_dep pat_info per_info args body gls = match args with @@ -1222,7 +1222,7 @@ let hrec_for fix_id per_info gls obj_id = try List.for_all2 eq_constr params per_info.per_params with Invalid_argument _ -> false end; let hd2 = applist (mkVar fix_id,args@[obj]) in - compose_lam rc (Reductionops.whd_beta gls.sigma hd2) + compose_lam rc (Reductionops.whd_beta gls.sigma (EConstr.of_constr hd2)) let warn_missing_case = CWarnings.create ~name:"declmode-missing-case" ~category:"declmode" diff --git a/plugins/extraction/extraction.ml b/plugins/extraction/extraction.ml index 85cacecdb..6ca34036a 100644 --- a/plugins/extraction/extraction.ml +++ b/plugins/extraction/extraction.ml @@ -74,7 +74,7 @@ type flag = info * scheme Really important function. *) let rec flag_of_type env t : flag = - let t = whd_all env none t in + let t = whd_all env none (EConstr.of_constr t) in match kind_of_term t with | Prod (x,t,c) -> flag_of_type (push_rel (LocalAssum (x,t)) env) c | Sort s when Sorts.is_prop s -> (Logic,TypeScheme) @@ -102,14 +102,14 @@ let is_info_scheme env t = match flag_of_type env t with (*s [type_sign] gernerates a signature aimed at treating a type application. *) let rec type_sign env c = - match kind_of_term (whd_all env none c) with + match kind_of_term (whd_all env none (EConstr.of_constr c)) with | Prod (n,t,d) -> (if is_info_scheme env t then Keep else Kill Kprop) :: (type_sign (push_rel_assum (n,t) env) d) | _ -> [] let rec type_scheme_nb_args env c = - match kind_of_term (whd_all env none c) with + match kind_of_term (whd_all env none (EConstr.of_constr c)) with | Prod (n,t,d) -> let n = type_scheme_nb_args (push_rel_assum (n,t) env) d in if is_info_scheme env t then n+1 else n @@ -135,7 +135,7 @@ let make_typvar n vl = next_ident_away id' vl let rec type_sign_vl env c = - match kind_of_term (whd_all env none c) with + match kind_of_term (whd_all env none (EConstr.of_constr c)) with | Prod (n,t,d) -> let s,vl = type_sign_vl (push_rel_assum (n,t) env) d in if not (is_info_scheme env t) then Kill Kprop::s, vl @@ -143,7 +143,7 @@ let rec type_sign_vl env c = | _ -> [],[] let rec nb_default_params env c = - match kind_of_term (whd_all env none c) with + match kind_of_term (whd_all env none (EConstr.of_constr c)) with | Prod (n,t,d) -> let n = nb_default_params (push_rel_assum (n,t) env) d in if is_default env t then n+1 else n @@ -214,7 +214,7 @@ let parse_ind_args si args relmax = let rec extract_type env db j c args = - match kind_of_term (whd_betaiotazeta Evd.empty c) with + match kind_of_term (whd_betaiotazeta none (EConstr.of_constr c)) with | App (d, args') -> (* We just accumulate the arguments. *) extract_type env db j d (Array.to_list args' @ args) @@ -297,7 +297,7 @@ and extract_type_app env db (r,s) args = let ml_args = List.fold_right (fun (b,c) a -> if b == Keep then - let p = List.length (fst (splay_prod env none (type_of env c))) in + let p = List.length (fst (splay_prod env none (EConstr.of_constr (type_of env c)))) in let db = iterate (fun l -> 0 :: l) p db in (extract_type_scheme env db c p) :: a else a) @@ -316,12 +316,12 @@ and extract_type_app env db (r,s) args = and extract_type_scheme env db c p = if Int.equal p 0 then extract_type env db 0 c [] else - let c = whd_betaiotazeta Evd.empty c in + let c = whd_betaiotazeta none (EConstr.of_constr c) in match kind_of_term c with | Lambda (n,t,d) -> extract_type_scheme (push_rel_assum (n,t) env) db d (p-1) | _ -> - let rels = fst (splay_prod env none (type_of env c)) in + let rels = fst (splay_prod env none (EConstr.of_constr (type_of env c))) in let env = push_rels_assum rels env in let eta_args = List.rev_map mkRel (List.interval 1 p) in extract_type env db 0 (lift p c) eta_args @@ -488,7 +488,7 @@ and extract_really_ind env kn mib = *) and extract_type_cons env db dbmap c i = - match kind_of_term (whd_all env none c) with + match kind_of_term (whd_all env none (EConstr.of_constr c)) with | Prod (n,t,d) -> let env' = push_rel_assum (n,t) env in let db' = (try Int.Map.find i dbmap with Not_found -> 0) :: db in @@ -846,7 +846,7 @@ and extract_fix env mle i (fi,ti,ci as recd) mlt = and decompose the term [c] in [n] lambdas, with eta-expansion if needed. *) let decomp_lams_eta_n n m env c t = - let rels = fst (splay_prod_n env none n t) in + let rels = fst (splay_prod_n env none n (EConstr.of_constr t)) in let rels = List.map (fun (LocalAssum (id,c) | LocalDef (id,_,c)) -> (id,c)) rels in let rels',c = decompose_lam c in let d = n - m in diff --git a/plugins/firstorder/instances.ml b/plugins/firstorder/instances.ml index eebd974ea..a3513692c 100644 --- a/plugins/firstorder/instances.ml +++ b/plugins/firstorder/instances.ml @@ -107,7 +107,7 @@ let mk_open_instance id idc gl m t= let typ=pf_unsafe_type_of gl idc in (* since we know we will get a product, reduction is not too expensive *) - let (nam,_,_)=destProd (whd_all env evmap typ) in + let (nam,_,_)=destProd (whd_all env evmap (EConstr.of_constr typ)) in match nam with Name id -> id | Anonymous -> dummy_bvid in diff --git a/plugins/firstorder/unify.ml b/plugins/firstorder/unify.ml index 01c019744..fb237f29b 100644 --- a/plugins/firstorder/unify.ml +++ b/plugins/firstorder/unify.ml @@ -41,8 +41,8 @@ let unif t1 t2= Queue.add (t1,t2) bige; try while true do let t1,t2=Queue.take bige in - let nt1=head_reduce (whd_betaiotazeta Evd.empty t1) - and nt2=head_reduce (whd_betaiotazeta Evd.empty t2) in + let nt1=head_reduce (whd_betaiotazeta Evd.empty (EConstr.of_constr t1)) + and nt2=head_reduce (whd_betaiotazeta Evd.empty (EConstr.of_constr t2)) in match (kind_of_term nt1),(kind_of_term nt2) with Meta i,Meta j-> if not (Int.equal i j) then diff --git a/plugins/funind/functional_principles_proofs.ml b/plugins/funind/functional_principles_proofs.ml index f6567ab81..258ee5ad6 100644 --- a/plugins/funind/functional_principles_proofs.ml +++ b/plugins/funind/functional_principles_proofs.ml @@ -318,7 +318,7 @@ let change_eq env sigma hyp_id (context:Context.Rel.t) x t end_of_type = context in let new_type_of_hyp = - Reductionops.nf_betaiota Evd.empty new_type_of_hyp in + Reductionops.nf_betaiota Evd.empty (EConstr.of_constr new_type_of_hyp) in let new_ctxt,new_end_of_type = decompose_prod_n_assum ctxt_size new_type_of_hyp in @@ -786,7 +786,7 @@ let build_proof do_finalize dyn_infos g | Lambda _ -> let new_term = - Reductionops.nf_beta Evd.empty dyn_infos.info in + Reductionops.nf_beta Evd.empty (EConstr.of_constr dyn_infos.info) in build_proof do_finalize {dyn_infos with info = new_term} g | LetIn _ -> @@ -1090,7 +1090,7 @@ let prove_princ_for_struct (evd:Evd.evar_map ref) interactive_proof fun_num fnam (CClosure.RedFlags.mkflags [CClosure.RedFlags.fZETA]) (Global.env ()) (Evd.empty) - body + (EConstr.of_constr body) | None -> error ( "Cannot define a principle over an axiom ") in let fbody = get_body fnames.(fun_num) in @@ -1142,8 +1142,8 @@ let prove_princ_for_struct (evd:Evd.evar_map ref) interactive_proof fun_num fnam Array.map (fun body -> Reductionops.nf_betaiota Evd.empty - (applist(substl (List.rev (Array.to_list all_funs_with_full_params)) body, - List.rev_map var_of_decl princ_params)) + (EConstr.of_constr (applist(substl (List.rev (Array.to_list all_funs_with_full_params)) body, + List.rev_map var_of_decl princ_params))) ) bodies in @@ -1179,20 +1179,20 @@ let prove_princ_for_struct (evd:Evd.evar_map ref) interactive_proof fun_num fnam let body_with_param,num = let body = get_body fnames.(i) in let body_with_full_params = - Reductionops.nf_betaiota Evd.empty ( - applist(body,List.rev_map var_of_decl full_params)) + Reductionops.nf_betaiota Evd.empty (EConstr.of_constr ( + applist(body,List.rev_map var_of_decl full_params))) in match kind_of_term body_with_full_params with | Fix((_,num),(_,_,bs)) -> Reductionops.nf_betaiota Evd.empty - ( + (EConstr.of_constr ( (applist (substl (List.rev (Array.to_list all_funs_with_full_params)) bs.(num), List.rev_map var_of_decl princ_params)) - ),num + )),num | _ -> error "Not a mutual block" in let info = @@ -1269,7 +1269,7 @@ let prove_princ_for_struct (evd:Evd.evar_map ref) interactive_proof fun_num fnam rec_hyps = []; info = Reductionops.nf_betaiota Evd.empty - (applist(fix_body,List.rev_map mkVar args_id)); + (EConstr.of_constr (applist(fix_body,List.rev_map mkVar args_id))); eq_hyps = [] } in @@ -1329,10 +1329,10 @@ let prove_princ_for_struct (evd:Evd.evar_map ref) interactive_proof fun_num fnam rec_hyps = []; info = Reductionops.nf_betaiota Evd.empty - (applist(fbody_with_full_params, + (EConstr.of_constr (applist(fbody_with_full_params, (List.rev_map var_of_decl princ_params)@ (List.rev_map mkVar args_id) - )); + ))); eq_hyps = [] } in diff --git a/plugins/funind/functional_principles_types.ml b/plugins/funind/functional_principles_types.ml index 032d887de..9637632a6 100644 --- a/plugins/funind/functional_principles_types.ml +++ b/plugins/funind/functional_principles_types.ml @@ -405,7 +405,7 @@ let get_funs_constant mp dp = (CClosure.RedFlags.mkflags [CClosure.RedFlags.fZETA]) (Global.env ()) (Evd.from_env (Global.env ())) - body + (EConstr.of_constr body) in body | None -> error ( "Cannot define a principle over an axiom ") diff --git a/plugins/funind/glob_term_to_relation.ml b/plugins/funind/glob_term_to_relation.ml index de2e5ea4e..92de4d873 100644 --- a/plugins/funind/glob_term_to_relation.ml +++ b/plugins/funind/glob_term_to_relation.ml @@ -352,7 +352,7 @@ let add_pat_variables pat typ env : Environ.env = | PatVar(_,na) -> Environ.push_rel (RelDecl.LocalAssum (na,typ)) env | PatCstr(_,c,patl,na) -> let Inductiveops.IndType(indf,indargs) = - try Inductiveops.find_rectype env (Evd.from_env env) typ + try Inductiveops.find_rectype env (Evd.from_env env) (EConstr.of_constr typ) with Not_found -> assert false in let constructors = Inductiveops.get_constructors env indf in @@ -409,7 +409,7 @@ let rec pattern_to_term_and_type env typ = function constr in let Inductiveops.IndType(indf,indargs) = - try Inductiveops.find_rectype env (Evd.from_env env) typ + try Inductiveops.find_rectype env (Evd.from_env env) (EConstr.of_constr typ) with Not_found -> assert false in let constructors = Inductiveops.get_constructors env indf in @@ -629,7 +629,7 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return = let b_as_constr,ctx = Pretyping.understand env (Evd.from_env env) b in let b_typ = Typing.unsafe_type_of env (Evd.from_env env) b_as_constr in let (ind,_) = - try Inductiveops.find_inductive env (Evd.from_env env) b_typ + try Inductiveops.find_inductive env (Evd.from_env env) (EConstr.of_constr b_typ) with Not_found -> user_err (str "Cannot find the inductive associated to " ++ Printer.pr_glob_constr b ++ str " in " ++ @@ -661,7 +661,7 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return = let b_as_constr,ctx = Pretyping.understand env (Evd.from_env env) b in let b_typ = Typing.unsafe_type_of env (Evd.from_env env) b_as_constr in let (ind,_) = - try Inductiveops.find_inductive env (Evd.from_env env) b_typ + try Inductiveops.find_inductive env (Evd.from_env env) (EConstr.of_constr b_typ) with Not_found -> user_err (str "Cannot find the inductive associated to " ++ Printer.pr_glob_constr b ++ str " in " ++ diff --git a/plugins/funind/invfun.ml b/plugins/funind/invfun.ml index cf42a809d..9abe60402 100644 --- a/plugins/funind/invfun.ml +++ b/plugins/funind/invfun.ml @@ -251,7 +251,7 @@ let prove_fun_correct evd functional_induction funs_constr graphs_constr schemes let mib,_ = Global.lookup_inductive graph_ind in (* and the principle to use in this lemma in $\zeta$ normal form *) let f_principle,princ_type = schemes.(i) in - let princ_type = nf_zeta princ_type in + let princ_type = nf_zeta (EConstr.of_constr princ_type) in let princ_infos = Tactics.compute_elim_sig princ_type in (* The number of args of the function is then easily computable *) let nb_fun_args = nb_prod (project g) (EConstr.of_constr (pf_concl g)) - 2 in @@ -428,7 +428,7 @@ let prove_fun_correct evd functional_induction funs_constr graphs_constr schemes List.rev (fst (List.fold_left2 (fun (bindings,avoid) decl p -> let id = Namegen.next_ident_away (Nameops.out_name (RelDecl.get_name decl)) avoid in - (nf_zeta p)::bindings,id::avoid) + (nf_zeta (EConstr.of_constr p))::bindings,id::avoid) ([],avoid) princ_infos.predicates (lemmas))) @@ -496,7 +496,7 @@ and intros_with_rewrite_aux : tactic = begin match kind_of_term t with | App(eq,args) when (eq_constr eq eq_ind) -> - if Reductionops.is_conv (pf_env g) (project g) args.(1) args.(2) + if Reductionops.is_conv (pf_env g) (project g) (EConstr.of_constr args.(1)) (EConstr.of_constr args.(2)) then let id = pf_get_new_id (Id.of_string "y") g in tclTHENSEQ [ Proofview.V82.of_tactic (Simple.intro id); thin [id]; intros_with_rewrite ] g @@ -655,12 +655,12 @@ let prove_fun_complete funcs graphs schemes lemmas_types_infos i : tactic = *) let lemmas = Array.map - (fun (_,(ctxt,concl)) -> nf_zeta (Termops.it_mkLambda_or_LetIn concl ctxt)) + (fun (_,(ctxt,concl)) -> nf_zeta (EConstr.of_constr (Termops.it_mkLambda_or_LetIn concl ctxt))) lemmas_types_infos in (* We get the constant and the principle corresponding to this lemma *) let f = funcs.(i) in - let graph_principle = nf_zeta schemes.(i) in + let graph_principle = nf_zeta (EConstr.of_constr schemes.(i)) in let princ_type = pf_unsafe_type_of g graph_principle in let princ_infos = Tactics.compute_elim_sig princ_type in (* Then we get the number of argument of the function @@ -793,7 +793,7 @@ let derive_correctness make_scheme functional_induction (funs: pconstant list) ( graphs_constr.(i) <- graph; let type_of_lemma = Termops.it_mkProd_or_LetIn type_of_lemma_concl type_of_lemma_ctxt in let _ = Typing.e_type_of (Global.env ()) evd type_of_lemma in - let type_of_lemma = nf_zeta type_of_lemma in + let type_of_lemma = nf_zeta (EConstr.of_constr type_of_lemma) in observe (str "type_of_lemma := " ++ Printer.pr_lconstr_env (Global.env ()) !evd type_of_lemma); type_of_lemma,type_info ) @@ -860,7 +860,7 @@ let derive_correctness make_scheme functional_induction (funs: pconstant list) ( let type_of_lemma = Termops.it_mkProd_or_LetIn type_of_lemma_concl type_of_lemma_ctxt in - let type_of_lemma = nf_zeta type_of_lemma in + let type_of_lemma = nf_zeta (EConstr.of_constr type_of_lemma) in observe (str "type_of_lemma := " ++ Printer.pr_lconstr type_of_lemma); type_of_lemma,type_info ) diff --git a/plugins/funind/merge.ml b/plugins/funind/merge.ml index 865042afb..222c0c804 100644 --- a/plugins/funind/merge.ml +++ b/plugins/funind/merge.ml @@ -135,7 +135,7 @@ let prNamedRLDecl s lc = let showind (id:Id.t) = let cstrid = Constrintern.global_reference id in - let ind1,cstrlist = Inductiveops.find_inductive (Global.env()) Evd.empty cstrid in + let ind1,cstrlist = Inductiveops.find_inductive (Global.env()) Evd.empty (EConstr.of_constr cstrid) in let mib1,ib1 = Inductive.lookup_mind_specif (Global.env()) (fst ind1) in List.iter (fun decl -> print_string (string_of_name (Context.Rel.Declaration.get_name decl) ^ ":"); diff --git a/plugins/funind/recdef.ml b/plugins/funind/recdef.ml index 6b63d7771..4fd9e0ff8 100644 --- a/plugins/funind/recdef.ml +++ b/plugins/funind/recdef.ml @@ -693,7 +693,7 @@ let mkDestructEq : (fun g2 -> let changefun patvars = { run = fun sigma -> let redfun = pattern_occs [Locus.AllOccurrencesBut [1], expr] in - redfun.Reductionops.e_redfun (pf_env g2) sigma (pf_concl g2) + redfun.Reductionops.e_redfun (pf_env g2) sigma (EConstr.of_constr (pf_concl g2)) } in Proofview.V82.of_tactic (change_in_concl None changefun) g2); Proofview.V82.of_tactic (simplest_case expr)]), to_revert @@ -1499,7 +1499,7 @@ let recursive_definition is_mes function_name rec_impls type_of_f r rec_arg_num (* Pp.msgnl (str "lemma type := " ++ Printer.pr_lconstr equation_lemma_type ++ fnl ()); *) let res_vars,eq' = decompose_prod equation_lemma_type in let env_eq' = Environ.push_rel_context (List.map (fun (x,y) -> LocalAssum (x,y)) res_vars) env in - let eq' = nf_zeta env_eq' eq' in + let eq' = nf_zeta env_eq' (EConstr.of_constr eq') in let res = (* Pp.msgnl (str "res_var :=" ++ Printer.pr_lconstr_env (push_rel_context (List.map (function (x,t) -> (x,None,t)) res_vars) env) eq'); *) (* Pp.msgnl (str "rec_arg_num := " ++ str (string_of_int rec_arg_num)); *) diff --git a/plugins/micromega/coq_micromega.ml b/plugins/micromega/coq_micromega.ml index 49fcf83b4..9fb1463db 100644 --- a/plugins/micromega/coq_micromega.ml +++ b/plugins/micromega/coq_micromega.ml @@ -902,7 +902,7 @@ struct let is_convertible gl t1 t2 = - Reductionops.is_conv (Tacmach.pf_env gl) (Tacmach.project gl) t1 t2 + Reductionops.is_conv (Tacmach.pf_env gl) (Tacmach.project gl) (EConstr.of_constr t1) (EConstr.of_constr t2) let parse_zop gl (op,args) = match kind_of_term op with diff --git a/plugins/omega/coq_omega.ml b/plugins/omega/coq_omega.ml index 1afc6500b..d15449aef 100644 --- a/plugins/omega/coq_omega.ml +++ b/plugins/omega/coq_omega.ml @@ -1591,7 +1591,7 @@ let nat_inject = (loop lit) ] | Kapp(Eq,[typ;t1;t2]) -> - if is_conv typ (Lazy.force coq_nat) then + if is_conv (EConstr.of_constr typ) (EConstr.of_constr (Lazy.force coq_nat)) then Tacticals.New.tclTHENLIST [ (generalize_tac [mkApp (Lazy.force coq_inj_eq, [| t1;t2;mkVar i |]) ]); diff --git a/plugins/quote/quote.ml b/plugins/quote/quote.ml index c6376727a..afc7e6665 100644 --- a/plugins/quote/quote.ml +++ b/plugins/quote/quote.ml @@ -249,8 +249,8 @@ let compute_ivs f cs gl = (* Then we test if the RHS is the RHS for variables *) else begin match decompose_app bodyi with | vmf, [_; _; a3; a4 ] - when isRel a3 && isRel a4 && is_conv vmf - (Lazy.force coq_varmap_find)-> + when isRel a3 && isRel a4 && is_conv (EConstr.of_constr vmf) + (EConstr.of_constr (Lazy.force coq_varmap_find)) -> v_lhs := Some (compute_lhs (snd (List.hd args3)) i nargsi) diff --git a/plugins/setoid_ring/newring.ml b/plugins/setoid_ring/newring.ml index 657efe175..cf0f51911 100644 --- a/plugins/setoid_ring/newring.ml +++ b/plugins/setoid_ring/newring.ml @@ -82,6 +82,7 @@ let lookup_map map = user_err ~hdr:"lookup_map" (str"map "++qs map++str"not found") let protect_red map env sigma c = + let c = EConstr.Unsafe.to_constr c in kl (create_clos_infos all env) (mk_clos_but (lookup_map map c) (Esubst.subs_id 0) c);; @@ -347,7 +348,7 @@ let find_ring_structure env sigma l = let ty = Retyping.get_type_of env sigma t in let check c = let ty' = Retyping.get_type_of env sigma c in - if not (Reductionops.is_conv env sigma ty ty') then + if not (Reductionops.is_conv env sigma (EConstr.of_constr ty) (EConstr.of_constr ty')) then user_err ~hdr:"ring" (str"arguments of ring_simplify do not have all the same type") in @@ -827,7 +828,7 @@ let find_field_structure env sigma l = let ty = Retyping.get_type_of env sigma t in let check c = let ty' = Retyping.get_type_of env sigma c in - if not (Reductionops.is_conv env sigma ty ty') then + if not (Reductionops.is_conv env sigma (EConstr.of_constr ty) (EConstr.of_constr ty')) then user_err ~hdr:"field" (str"arguments of field_simplify do not have all the same type") in diff --git a/plugins/ssrmatching/ssrmatching.ml4 b/plugins/ssrmatching/ssrmatching.ml4 index 86cc928c8..cc39b7260 100644 --- a/plugins/ssrmatching/ssrmatching.ml4 +++ b/plugins/ssrmatching/ssrmatching.ml4 @@ -481,7 +481,9 @@ let evars_for_FO ~hack env sigma0 (ise0:evar_map) c0 = (* p_origin can be passed to obtain a better error message *) let mk_tpattern ?p_origin ?(hack=false) env sigma0 (ise, t) ok dir p = let k, f, a = - let f, a = Reductionops.whd_betaiota_stack ise p in + let f, a = Reductionops.whd_betaiota_stack ise (EConstr.of_constr p) in + let f = EConstr.Unsafe.to_constr f in + let a = List.map EConstr.Unsafe.to_constr a in match kind_of_term f with | Const (p,_) -> let np = proj_nparams p in diff --git a/pretyping/cases.ml b/pretyping/cases.ml index 6b480986c..be72091a9 100644 --- a/pretyping/cases.ml +++ b/pretyping/cases.ml @@ -292,7 +292,7 @@ let inductive_template evdref env tmloc ind = applist (mkIndU indu,List.rev evarl) let try_find_ind env sigma typ realnames = - let (IndType(indf,realargs) as ind) = find_rectype env sigma typ in + let (IndType(indf,realargs) as ind) = find_rectype env sigma (EConstr.of_constr typ) in let names = match realnames with | Some names -> names @@ -1035,7 +1035,7 @@ let specialize_predicate newtomatchs (names,depna) arsign cs tms ccl = (* We need _parallel_ bindings to get gamma, x1...xn |- PI tms. ccl'' *) (* Note: applying the substitution in tms is not important (is it sure?) *) let ccl'' = - whd_betaiota Evd.empty (subst_predicate (realargsi, copti) ccl' tms) in + whd_betaiota Evd.empty (EConstr.of_constr (subst_predicate (realargsi, copti) ccl' tms)) in (* We adjust ccl st: gamma, x'1..x'n, x1..xn, tms |- ccl'' *) let ccl''' = liftn_predicate n (n+1) ccl'' tms in (* We finally get gamma,x'1..x'n,x |- [X1;x1:I(X1)]..[Xn;xn:I(Xn)]pred'''*) @@ -1044,7 +1044,7 @@ let specialize_predicate newtomatchs (names,depna) arsign cs tms ccl = let find_predicate loc env evdref p current (IndType (indf,realargs)) dep tms = let pred = abstract_predicate env !evdref indf current realargs dep tms p in (pred, whd_betaiota !evdref - (applist (pred, realargs@[current]))) + (EConstr.of_constr (applist (pred, realargs@[current])))) (* Take into account that a type has been discovered to be inductive, leading to more dependencies in the predicate if the type has indices *) @@ -1199,7 +1199,7 @@ let rec generalize_problem names pb = function | LocalDef (Anonymous,_,_) -> pb', deps | _ -> (* for better rendering *) - let d = RelDecl.map_type (whd_betaiota !(pb.evdref)) d in + let d = RelDecl.map_type (fun c -> whd_betaiota !(pb.evdref) (EConstr.of_constr c)) d in let tomatch = lift_tomatch_stack 1 pb'.tomatch in let tomatch = relocate_index_tomatch (i+1) 1 tomatch in { pb' with @@ -1377,7 +1377,7 @@ and match_current pb (initial,tomatch) = find_predicate pb.caseloc pb.env pb.evdref pred current indt (names,dep) tomatch in let ci = make_case_info pb.env (fst mind) pb.casestyle in - let pred = nf_betaiota !(pb.evdref) pred in + let pred = nf_betaiota !(pb.evdref) (EConstr.of_constr pred) in let case = make_case_or_project pb.env indf ci pred current brvals in @@ -1613,7 +1613,7 @@ let rec list_assoc_in_triple x = function *) let abstract_tycon loc env evdref subst tycon extenv t = - let t = nf_betaiota !evdref t in (* it helps in some cases to remove K-redex*) + let t = nf_betaiota !evdref (EConstr.of_constr t) in (* it helps in some cases to remove K-redex*) let src = match kind_of_term t with | Evar (evk,_) -> (loc,Evar_kinds.SubEvar evk) | _ -> (loc,Evar_kinds.CasesType true) in @@ -1635,13 +1635,13 @@ let abstract_tycon loc env evdref subst tycon extenv t = try list_assoc_in_triple i subst0 with Not_found -> mkRel i) 1 (rel_context env) in let ev' = e_new_evar env evdref ~src ty in - begin match solve_simple_eqn (evar_conv_x full_transparent_state) env !evdref (None,ev,substl inst ev') with + begin match solve_simple_eqn (evar_conv_x full_transparent_state) env !evdref (None,ev,EConstr.of_constr (substl inst ev')) with | Success evd -> evdref := evd | UnifFailure _ -> assert false end; ev' | _ -> - let good = List.filter (fun (_,u,_) -> is_conv_leq env !evdref t u) subst in + let good = List.filter (fun (_,u,_) -> is_conv_leq env !evdref (EConstr.of_constr t) (EConstr.of_constr u)) subst in match good with | [] -> let self env c = EConstr.of_constr (aux env (EConstr.Unsafe.to_constr c)) in @@ -1705,7 +1705,7 @@ let build_inversion_problem loc env sigma tms t = let id = next_name_away (named_hd env t Anonymous) avoid in PatVar (Loc.ghost,Name id), ((id,t)::subst, id::avoid) in let rec reveal_pattern t (subst,avoid as acc) = - match kind_of_term (whd_all env sigma t) with + match kind_of_term (whd_all env sigma (EConstr.of_constr t)) with | Construct (cstr,u) -> PatCstr (Loc.ghost,cstr,[],Anonymous), acc | App (f,v) when isConstruct f -> let cstr,u = destConstruct f in @@ -2038,7 +2038,7 @@ let constr_of_pat env evdref arsign pat avoid = | PatCstr (l,((_, i) as cstr),args,alias) -> let cind = inductive_of_constructor cstr in let IndType (indf, _) = - try find_rectype env ( !evdref) (lift (-(List.length realargs)) ty) + try find_rectype env ( !evdref) (EConstr.of_constr (lift (-(List.length realargs)) ty)) with Not_found -> error_case_not_inductive env !evdref {uj_val = ty; uj_type = Typing.unsafe_type_of env !evdref ty} in @@ -2068,7 +2068,7 @@ let constr_of_pat env evdref arsign pat avoid = let app = applistc cstr (List.map (lift (List.length sign)) params) in let app = applistc app args in let apptype = Retyping.get_type_of env ( !evdref) app in - let IndType (indf, realargs) = find_rectype env ( !evdref) apptype in + let IndType (indf, realargs) = find_rectype env (!evdref) (EConstr.of_constr apptype) in match alias with Anonymous -> pat', sign, app, apptype, realargs, n, avoid @@ -2327,7 +2327,7 @@ let build_dependent_signature env evdref avoid tomatchs arsign = let t = RelDecl.get_type decl in let argt = Retyping.get_type_of env !evdref arg in let eq, refl_arg = - if Reductionops.is_conv env !evdref argt t then + if Reductionops.is_conv env !evdref (EConstr.of_constr argt) (EConstr.of_constr t) then (mk_eq evdref (lift (nargeqs + slift) argt) (mkRel (nargeqs + slift)) (lift (nargeqs + nar) arg), diff --git a/pretyping/classops.ml b/pretyping/classops.ml index 30d100af9..fd21f5bd1 100644 --- a/pretyping/classops.ml +++ b/pretyping/classops.ml @@ -192,13 +192,14 @@ let coercion_exists coe = CoeTypMap.mem coe !coercion_tab (* find_class_type : evar_map -> constr -> cl_typ * universe_list * constr list *) let find_class_type sigma t = - let t', args = Reductionops.whd_betaiotazeta_stack sigma t in - match kind_of_term t' with - | Var id -> CL_SECVAR id, Univ.Instance.empty, args - | Const (sp,u) -> CL_CONST sp, u, args + let inj = EConstr.Unsafe.to_constr in + let t', args = Reductionops.whd_betaiotazeta_stack sigma (EConstr.of_constr t) in + match EConstr.kind sigma t' with + | Var id -> CL_SECVAR id, Univ.Instance.empty, List.map inj args + | Const (sp,u) -> CL_CONST sp, u, List.map inj args | Proj (p, c) when not (Projection.unfolded p) -> - CL_PROJ (Projection.constant p), Univ.Instance.empty, c :: args - | Ind (ind_sp,u) -> CL_IND ind_sp, u, args + CL_PROJ (Projection.constant p), Univ.Instance.empty, List.map inj (c :: args) + | Ind (ind_sp,u) -> CL_IND ind_sp, u, List.map inj args | Prod (_,_,_) -> CL_FUN, Univ.Instance.empty, [] | Sort _ -> CL_SORT, Univ.Instance.empty, [] | _ -> raise Not_found @@ -232,7 +233,7 @@ let class_of env sigma t = let (i, { cl_param = n1 } ) = class_info cl in (t, n1, i, u, args) with Not_found -> - let t = Tacred.hnf_constr env sigma t in + let t = Tacred.hnf_constr env sigma (EConstr.of_constr t) in let (cl, u, args) = find_class_type sigma t in let (i, { cl_param = n1 } ) = class_info cl in (t, n1, i, u, args) @@ -276,7 +277,7 @@ let apply_on_class_of env sigma t cont = t, cont i with Not_found -> (* Is it worth to be more incremental on the delta steps? *) - let t = Tacred.hnf_constr env sigma t in + let t = Tacred.hnf_constr env sigma (EConstr.of_constr t) in let (cl, u, args) = find_class_type sigma t in let (i, { cl_param = n1 } ) = class_info cl in if not (Int.equal (List.length args) n1) then raise Not_found; @@ -297,9 +298,9 @@ let lookup_path_to_sort_from env sigma s = let get_coercion_constructor env coe = let c, _ = - Reductionops.whd_all_stack env Evd.empty coe.coe_value + Reductionops.whd_all_stack env Evd.empty (EConstr.of_constr coe.coe_value) in - match kind_of_term c with + match EConstr.kind Evd.empty (** FIXME *) c with | Construct (cstr,u) -> (cstr, Inductiveops.constructor_nrealargs cstr -1) | _ -> @@ -403,7 +404,7 @@ type coercion = { let reference_arity_length ref = let t = Universes.unsafe_type_of_global ref in - List.length (fst (Reductionops.splay_arity (Global.env()) Evd.empty t)) + List.length (fst (Reductionops.splay_arity (Global.env()) Evd.empty (EConstr.of_constr t))) (** FIXME *) let projection_arity_length p = let len = reference_arity_length (ConstRef p) in diff --git a/pretyping/coercion.ml b/pretyping/coercion.ml index a3970fc0f..b062da1f4 100644 --- a/pretyping/coercion.ml +++ b/pretyping/coercion.ml @@ -63,7 +63,7 @@ let apply_coercion_args env evd check isproj argl funj = { uj_val = applist (j_val funj,argl); uj_type = typ } | h::restl -> (* On devrait pouvoir s'arranger pour qu'on n'ait pas a faire hnf_constr *) - match kind_of_term (whd_all env evd typ) with + match kind_of_term (whd_all env evd (EConstr.of_constr typ)) with | Prod (_,c1,c2) -> if check && not (e_cumul env evdref (Retyping.get_type_of env evd h) c1) then raise NoCoercion; @@ -95,7 +95,7 @@ let make_existential loc ?(opaque = not (get_proofs_transparency ())) env evdre Evarutil.e_new_evar env evdref ~src c let app_opt env evdref f t = - whd_betaiota !evdref (app_opt f t) + whd_betaiota !evdref (EConstr.of_constr (app_opt f t)) let pair_of_array a = (a.(0), a.(1)) @@ -128,11 +128,11 @@ let lift_args n sign = let mu env evdref t = let rec aux v = - let v' = hnf env !evdref v in + let v' = hnf env !evdref (EConstr.of_constr v) in match disc_subset v' with | Some (u, p) -> let f, ct = aux u in - let p = hnf_nodelta env !evdref p in + let p = hnf_nodelta env !evdref (EConstr.of_constr p) in (Some (fun x -> app_opt env evdref f (papp evdref sig_proj1 [| u; p; x |])), @@ -145,7 +145,7 @@ and coerce loc env evdref (x : Term.constr) (y : Term.constr) = let open Context.Rel.Declaration in let rec coerce_unify env x y = - let x = hnf env !evdref x and y = hnf env !evdref y in + let x = hnf env !evdref (EConstr.of_constr x) and y = hnf env !evdref (EConstr.of_constr y) in try evdref := the_conv_x_leq env x y !evdref; None @@ -153,7 +153,7 @@ and coerce loc env evdref (x : Term.constr) (y : Term.constr) and coerce' env x y : (Term.constr -> Term.constr) option = let subco () = subset_coerce env evdref x y in let dest_prod c = - match Reductionops.splay_prod_n env ( !evdref) 1 c with + match Reductionops.splay_prod_n env (!evdref) 1 (EConstr.of_constr c) with | [LocalAssum (na,t) | LocalDef (na,_,t)], c -> (na,t), c | _ -> raise NoSubtacCoercion in @@ -337,7 +337,7 @@ let app_coercion env evdref coercion v = | None -> v | Some f -> let v' = Typing.e_solve_evars env evdref (f v) in - whd_betaiota !evdref v' + whd_betaiota !evdref (EConstr.of_constr v') let coerce_itf loc env evd v t c1 = let evdref = ref evd in @@ -373,7 +373,7 @@ let apply_coercion env sigma p hj typ_cl = (* Try to coerce to a funclass; raise NoCoercion if not possible *) let inh_app_fun_core env evd j = - let t = whd_all env evd j.uj_type in + let t = whd_all env evd (EConstr.of_constr j.uj_type) in match kind_of_term t with | Prod (_,_,_) -> (evd,j) | Evar ev -> @@ -414,7 +414,7 @@ let inh_tosort_force loc env evd j = error_not_a_type ~loc env evd j let inh_coerce_to_sort loc env evd j = - let typ = whd_all env evd j.uj_type in + let typ = whd_all env evd (EConstr.of_constr j.uj_type) in match kind_of_term typ with | Sort s -> (evd,{ utj_val = j.uj_val; utj_type = s }) | Evar ev when not (is_defined evd (fst ev)) -> @@ -466,8 +466,8 @@ let rec inh_conv_coerce_to_fail loc env evd rigidonly v t c1 = try inh_coerce_to_fail env evd rigidonly v t c1 with NoCoercion -> match - kind_of_term (whd_all env evd t), - kind_of_term (whd_all env evd c1) + kind_of_term (whd_all env evd (EConstr.of_constr t)), + kind_of_term (whd_all env evd (EConstr.of_constr c1)) with | Prod (name,t1,t2), Prod (_,u1,u2) -> (* Conversion did not work, we may succeed with a coercion. *) @@ -485,7 +485,7 @@ let rec inh_conv_coerce_to_fail loc env evd rigidonly v t c1 = inh_conv_coerce_to_fail loc env1 evd rigidonly (Some (mkRel 1)) (lift 1 u1) (lift 1 t1) in let v1 = Option.get v1 in - let v2 = Option.map (fun v -> beta_applist (lift 1 v,[v1])) v in + let v2 = Option.map (fun v -> beta_applist evd' (EConstr.of_constr (lift 1 v),[EConstr.of_constr v1])) v in let t2 = match v2 with | None -> subst_term evd' (EConstr.of_constr v1) (EConstr.of_constr t2) | Some v2 -> Retyping.get_type_of env1 evd' v2 in diff --git a/pretyping/constr_matching.ml b/pretyping/constr_matching.ml index d7b73d333..66e690714 100644 --- a/pretyping/constr_matching.ml +++ b/pretyping/constr_matching.ml @@ -165,7 +165,7 @@ let matches_core env sigma convert allow_partial_app allow_bound_rels | _, _ -> (if convert then let sigma,c' = Evd.fresh_global env sigma ref in - is_conv env sigma c' c + is_conv env sigma (EConstr.of_constr c') (EConstr.of_constr c) else false) in let rec sorec ctx env subst p t = diff --git a/pretyping/detyping.ml b/pretyping/detyping.ml index 72cf31010..a4d943cfa 100644 --- a/pretyping/detyping.ml +++ b/pretyping/detyping.ml @@ -502,7 +502,7 @@ let rec detype flags avoid env sigma t = let pb = Environ.lookup_projection p (snd env) in let body = pb.Declarations.proj_body in let ty = Retyping.get_type_of (snd env) sigma c in - let ((ind,u), args) = Inductiveops.find_mrectype (snd env) sigma ty in + let ((ind,u), args) = Inductiveops.find_mrectype (snd env) sigma (EConstr.of_constr ty) in let body' = strip_lam_assum body in let body' = subst_instance_constr u body' in substl (c :: List.rev args) body' diff --git a/pretyping/evarconv.ml b/pretyping/evarconv.ml index 194d0b297..f54a57d57 100644 --- a/pretyping/evarconv.ml +++ b/pretyping/evarconv.ml @@ -78,6 +78,7 @@ type flex_kind_of_term = | Flexible of existential let flex_kind_of_term ts env evd c sk = + let c = EConstr.Unsafe.to_constr c in match kind_of_term c with | LetIn _ | Rel _ | Const _ | Var _ | Proj _ -> Option.cata (fun x -> MaybeFlexible x) Rigid (eval_flexible_term ts env evd c) @@ -88,10 +89,12 @@ let flex_kind_of_term ts env evd c sk = | Fix _ -> Rigid (* happens when the fixpoint is partially applied *) | Cast _ | App _ | Case _ -> assert false +let zip evd (c, stk) = EConstr.Unsafe.to_constr (Stack.zip evd (c, stk)) + let apprec_nohdbeta ts env evd c = - let (t,sk as appr) = Reductionops.whd_nored_state evd (c, []) in + let (t,sk as appr) = Reductionops.whd_nored_state evd (EConstr.of_constr c, []) in if Stack.not_purely_applicative sk - then Stack.zip (fst (whd_betaiota_deltazeta_for_iota_state + then zip evd (fst (whd_betaiota_deltazeta_for_iota_state ts env evd Cst_stack.empty appr)) else c @@ -135,6 +138,8 @@ let occur_rigidly (evk,_ as ev) evd t = projection would have been reduced) *) let check_conv_record env sigma (t1,sk1) (t2,sk2) = + let t1 = EConstr.Unsafe.to_constr t1 in + let t2 = EConstr.Unsafe.to_constr t2 in let (proji, u), arg = Universes.global_app_of_constr t1 in let canon_s,sk2_effective = try @@ -143,7 +148,7 @@ let check_conv_record env sigma (t1,sk1) (t2,sk2) = let _, a, b = destProd (Evarutil.nf_evar sigma t2) in if EConstr.Vars.noccurn sigma 1 (EConstr.of_constr b) then lookup_canonical_conversion (proji, Prod_cs), - (Stack.append_app [|a;pop (EConstr.of_constr b)|] Stack.empty) + (Stack.append_app [|EConstr.of_constr a;EConstr.of_constr (pop (EConstr.of_constr b))|] Stack.empty) else raise Not_found | Sort s -> lookup_canonical_conversion @@ -162,12 +167,12 @@ let check_conv_record env sigma (t1,sk1) (t2,sk2) = | Some c -> (* A primitive projection applied to c *) let ty = Retyping.get_type_of ~lax:true env sigma c in let (i,u), ind_args = - try Inductiveops.find_mrectype env sigma ty + try Inductiveops.find_mrectype env sigma (EConstr.of_constr ty) with _ -> raise Not_found - in Stack.append_app_list ind_args Stack.empty, c, sk1 + in Stack.append_app_list (List.map EConstr.of_constr ind_args) Stack.empty, c, sk1 | None -> match Stack.strip_n_app nparams sk1 with - | Some (params1, c1, extra_args1) -> params1, c1, extra_args1 + | Some (params1, c1, extra_args1) -> params1, EConstr.Unsafe.to_constr c1, extra_args1 | _ -> raise Not_found in let us2,extra_args2 = let l_us = List.length us in @@ -180,9 +185,9 @@ let check_conv_record env sigma (t1,sk1) (t2,sk2) = let t' = subst_univs_level_constr subst t' in let bs' = List.map (subst_univs_level_constr subst) bs in let h, _ = decompose_app_vect sigma (EConstr.of_constr t') in - ctx',(h, t2),c',bs',(Stack.append_app_list params Stack.empty,params1), - (Stack.append_app_list us Stack.empty,us2),(extra_args1,extra_args2),c1, - (n,Stack.zip(t2,sk2)) + ctx',(h, t2),c',bs',(Stack.append_app_list (List.map EConstr.of_constr params) Stack.empty,params1), + (Stack.append_app_list (List.map EConstr.of_constr us) Stack.empty,us2),(extra_args1,extra_args2),c1, + (n, zip sigma (EConstr.of_constr t2,sk2)) (* Precondition: one of the terms of the pb is an uninstantiated evar, * possibly applied to arguments. *) @@ -212,10 +217,11 @@ let ise_exact ise x1 x2 = | Some _, Success i -> UnifFailure (i,NotSameArgSize) let ise_array2 evd f v1 v2 = + let inj c = EConstr.Unsafe.to_constr c in let rec allrec i = function | -1 -> Success i | n -> - match f i v1.(n) v2.(n) with + match f i (inj v1.(n)) (inj v2.(n)) with | Success i' -> allrec i' (n-1) | UnifFailure _ as x -> x in let lv1 = Array.length v1 in @@ -225,28 +231,35 @@ let ise_array2 evd f v1 v2 = (* Applicative node of stack are read from the outermost to the innermost but are unified the other way. *) let rec ise_app_stack2 env f evd sk1 sk2 = + let inj = EConstr.Unsafe.to_constr in match sk1,sk2 with | Stack.App node1 :: q1, Stack.App node2 :: q2 -> let (t1,l1) = Stack.decomp_node_last node1 q1 in let (t2,l2) = Stack.decomp_node_last node2 q2 in begin match ise_app_stack2 env f evd l1 l2 with |(_,UnifFailure _) as x -> x - |x,Success i' -> x,f env i' CONV t1 t2 + |x,Success i' -> x,f env i' CONV (inj t1) (inj t2) end | _, _ -> (sk1,sk2), Success evd +let push_rec_types pfix env = + let (i, c, t) = pfix in + let inj c = EConstr.Unsafe.to_constr c in + push_rec_types (i, Array.map inj c, Array.map inj t) env + (* This function tries to unify 2 stacks element by element. It works from the end to the beginning. If it unifies a non empty suffix of stacks but not the entire stacks, the first part of the answer is Some(the remaining prefixes to tackle)) *) let ise_stack2 no_app env evd f sk1 sk2 = + let inj = EConstr.Unsafe.to_constr in let rec ise_stack2 deep i sk1 sk2 = let fail x = if deep then Some (List.rev sk1, List.rev sk2), Success i else None, x in match sk1, sk2 with | [], [] -> None, Success i | Stack.Case (_,t1,c1,_)::q1, Stack.Case (_,t2,c2,_)::q2 -> - (match f env i CONV t1 t2 with + (match f env i CONV (inj t1) (inj t2) with | Success i' -> (match ise_array2 i' (fun ii -> f env ii CONV) c1 c2 with | Success i'' -> ise_stack2 true i'' q1 q2 @@ -279,6 +292,7 @@ let ise_stack2 no_app env evd f sk1 sk2 = (* Make sure that the matching suffix is the all stack *) let exact_ise_stack2 env evd f sk1 sk2 = + let inj = EConstr.Unsafe.to_constr in let rec ise_stack2 i sk1 sk2 = match sk1, sk2 with | [], [] -> Success i @@ -286,7 +300,7 @@ let exact_ise_stack2 env evd f sk1 sk2 = ise_and i [ (fun i -> ise_stack2 i q1 q2); (fun i -> ise_array2 i (fun ii -> f env ii CONV) c1 c2); - (fun i -> f env i CONV t1 t2)] + (fun i -> f env i CONV (inj t1) (inj t2))] | Stack.Fix (((li1, i1),(_,tys1,bds1 as recdef1)),a1,_)::q1, Stack.Fix (((li2, i2),(_,tys2,bds2)),a2,_)::q2 -> if Int.equal i1 i2 && Array.equal Int.equal li1 li2 then @@ -344,19 +358,19 @@ let rec evar_conv_x ts env evd pbty term1 term2 = let term2 = apprec_nohdbeta (fst ts) env evd term2 in let default () = evar_eqappr_x ts env evd pbty - (whd_nored_state evd (term1,Stack.empty), Cst_stack.empty) - (whd_nored_state evd (term2,Stack.empty), Cst_stack.empty) + (whd_nored_state evd (EConstr.of_constr term1,Stack.empty), Cst_stack.empty) + (whd_nored_state evd (EConstr.of_constr term2,Stack.empty), Cst_stack.empty) in begin match kind_of_term term1, kind_of_term term2 with | Evar ev, _ when Evd.is_undefined evd (fst ev) -> (match solve_simple_eqn (evar_conv_x ts) env evd - (position_problem true pbty,ev,term2) with + (position_problem true pbty,ev, EConstr.of_constr term2) with | UnifFailure (_,OccurCheck _) -> (* Eta-expansion might apply *) default () | x -> x) | _, Evar ev when Evd.is_undefined evd (fst ev) -> (match solve_simple_eqn (evar_conv_x ts) env evd - (position_problem false pbty,ev,term1) with + (position_problem false pbty,ev, EConstr.of_constr term1) with | UnifFailure (_, OccurCheck _) -> (* Eta-expansion might apply *) default () | x -> x) @@ -369,23 +383,25 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty UnifFailure (i, NotSameHead) in let miller_pfenning on_left fallback ev lF tM evd = + let lF = List.map EConstr.Unsafe.to_constr lF in match is_unification_pattern_evar env evd ev lF tM with | None -> fallback () | Some l1' -> (* Miller-Pfenning's patterns unification *) let t2 = nf_evar evd tM in let t2 = solve_pattern_eqn env evd l1' t2 in solve_simple_eqn (evar_conv_x ts) env evd - (position_problem on_left pbty,ev,t2) + (position_problem on_left pbty,ev, EConstr.of_constr t2) in let consume_stack on_left (termF,skF) (termO,skO) evd = + let inj = EConstr.Unsafe.to_constr in let switch f a b = if on_left then f a b else f b a in let not_only_app = Stack.not_purely_applicative skO in match switch (ise_stack2 not_only_app env evd (evar_conv_x ts)) skF skO with |Some (l,r), Success i' when on_left && (not_only_app || List.is_empty l) -> - switch (evar_conv_x ts env i' pbty) (Stack.zip(termF,l)) (Stack.zip(termO,r)) + switch (evar_conv_x ts env i' pbty) (zip evd (termF,l)) (zip evd (termO,r)) |Some (r,l), Success i' when not on_left && (not_only_app || List.is_empty l) -> - switch (evar_conv_x ts env i' pbty) (Stack.zip(termF,l)) (Stack.zip(termO,r)) - |None, Success i' -> switch (evar_conv_x ts env i' pbty) termF termO + switch (evar_conv_x ts env i' pbty) (zip evd (termF,l)) (zip evd (termO,r)) + |None, Success i' -> switch (evar_conv_x ts env i' pbty) (inj termF) (inj termO) |_, (UnifFailure _ as x) -> x |Some _, _ -> UnifFailure (evd,NotSameArgSize) in let eta env evd onleft sk term sk' term' = @@ -394,15 +410,15 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty let c = nf_evar evd c1 in let env' = push_rel (RelDecl.LocalAssum (na,c)) env in let out1 = whd_betaiota_deltazeta_for_iota_state - (fst ts) env' evd Cst_stack.empty (c'1, Stack.empty) in + (fst ts) env' evd Cst_stack.empty (EConstr.of_constr c'1, Stack.empty) in let out2 = whd_nored_state evd - (Stack.zip (term', sk' @ [Stack.Shift 1]), Stack.append_app [|mkRel 1|] Stack.empty), + (Stack.zip evd (term', sk' @ [Stack.Shift 1]), Stack.append_app [|EConstr.mkRel 1|] Stack.empty), Cst_stack.empty in if onleft then evar_eqappr_x ts env' evd CONV out1 out2 else evar_eqappr_x ts env' evd CONV out2 out1 in let rigids env evd sk term sk' term' = - let univs = Universes.eq_constr_universes term term' in + let univs = EConstr.eq_constr_universes evd term term' in match univs with | Some univs -> ise_and evd [(fun i -> @@ -420,10 +436,10 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty match Stack.list_of_app_stack skF with | None -> quick_fail evd | Some lF -> - let tM = Stack.zip apprM in + let tM = zip evd apprM in miller_pfenning on_left (fun () -> if not_only_app then (* Postpone the use of an heuristic *) - switch (fun x y -> Success (add_conv_pb (pbty,env,x,y) i)) (Stack.zip apprF) tM + switch (fun x y -> Success (add_conv_pb (pbty,env,x,y) i)) (zip evd apprF) tM else quick_fail i) ev lF tM i and consume (termF,skF as apprF) (termM,skM as apprM) i = @@ -437,7 +453,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty in let default i = ise_try i [f1; consume apprF apprM; delta] in - match kind_of_term termM with + match EConstr.kind evd termM with | Proj (p, c) when not (Stack.is_empty skF) -> (* Might be ?X args = p.c args', and we have to eta-expand the primitive projection if |args| >= |args'|+1. *) @@ -448,10 +464,10 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty else let f = try - let termM' = Retyping.expand_projection env evd p c [] in + let termM' = Retyping.expand_projection env evd p (EConstr.Unsafe.to_constr c) [] in let apprM', cstsM' = whd_betaiota_deltazeta_for_iota_state - (fst ts) env evd cstsM (termM',skM) + (fst ts) env evd cstsM (EConstr.of_constr termM',skM) in let delta' i = switch (evar_eqappr_x ts env i pbty) (apprF,cstsF) (apprM',cstsM') @@ -467,9 +483,9 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty let flex_rigid on_left ev (termF, skF as apprF) (termR, skR as apprR) = let switch f a b = if on_left then f a b else f b a in let eta evd = - match kind_of_term termR with + match EConstr.kind evd termR with | Lambda _ when (* if ever problem is ill-typed: *) List.is_empty skR -> - eta env evd false skR termR skF termF + eta env evd false skR (EConstr.Unsafe.to_constr termR) skF termF | Construct u -> eta_constructor ts env evd skR u skF termF | _ -> UnifFailure (evd,NotSameHead) in @@ -477,7 +493,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty | None -> ise_try evd [consume_stack on_left apprF apprR; eta] | Some lF -> - let tR = Stack.zip apprR in + let tR = zip evd apprR in miller_pfenning on_left (fun () -> ise_try evd @@ -487,10 +503,10 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty let i,tF = if isRel tR || isVar tR then (* Optimization so as to generate candidates *) - let i,ev = evar_absorb_arguments env i ev lF in + let i,ev = evar_absorb_arguments env i ev (List.map EConstr.Unsafe.to_constr lF) in i,mkEvar ev else - i,Stack.zip apprF in + i,zip evd apprF in switch (fun x y -> Success (add_conv_pb (pbty,env,x,y) i)) tF tR else @@ -516,20 +532,20 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty let ev1' = whd_evar i' (mkEvar ev1) in if isEvar ev1' then solve_simple_eqn (evar_conv_x ts) env i' - (position_problem true pbty,destEvar ev1',term2) + (position_problem true pbty,destEvar ev1', term2) else evar_eqappr_x ts env evd pbty - ((ev1', sk1), csts1) ((term2, sk2), csts2) + ((EConstr.of_constr ev1', sk1), csts1) ((term2, sk2), csts2) | Some (r,[]), Success i' -> (* We have sk1'[] = sk2[] for some sk1' s.t. sk1[]=sk1'[r[]] *) (* we now unify r[?ev1] and ?ev2 *) let ev2' = whd_evar i' (mkEvar ev2) in if isEvar ev2' then solve_simple_eqn (evar_conv_x ts) env i' - (position_problem false pbty,destEvar ev2',Stack.zip(term1,r)) + (position_problem false pbty,destEvar ev2',Stack.zip evd (term1,r)) else evar_eqappr_x ts env evd pbty - ((ev2', sk1), csts1) ((term2, sk2), csts2) + ((EConstr.of_constr ev2', sk1), csts1) ((term2, sk2), csts2) | Some ([],r), Success i' -> (* Symmetrically *) (* We have sk1[] = sk2'[] for some sk2' s.t. sk2[]=sk2'[r[]] *) @@ -537,9 +553,9 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty let ev1' = whd_evar i' (mkEvar ev1) in if isEvar ev1' then solve_simple_eqn (evar_conv_x ts) env i' - (position_problem true pbty,destEvar ev1',Stack.zip(term2,r)) + (position_problem true pbty,destEvar ev1',Stack.zip evd (term2,r)) else evar_eqappr_x ts env evd pbty - ((ev1', sk1), csts1) ((term2, sk2), csts2) + ((EConstr.of_constr ev1', sk1), csts1) ((term2, sk2), csts2) | None, (UnifFailure _ as x) -> (* sk1 and sk2 have no common outer part *) if Stack.not_purely_applicative sk2 then @@ -584,13 +600,13 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty ise_try evd [f1; f2] | Flexible ev1, MaybeFlexible v2 -> - flex_maybeflex true ev1 (appr1,csts1) (appr2,csts2) v2 + flex_maybeflex true ev1 (appr1,csts1) (appr2,csts2) (EConstr.of_constr v2) | MaybeFlexible v1, Flexible ev2 -> - flex_maybeflex false ev2 (appr2,csts2) (appr1,csts1) v1 + flex_maybeflex false ev2 (appr2,csts2) (appr1,csts1) (EConstr.of_constr v1) | MaybeFlexible v1, MaybeFlexible v2 -> begin - match kind_of_term term1, kind_of_term term2 with + match kind_of_term (EConstr.Unsafe.to_constr term1), kind_of_term (EConstr.Unsafe.to_constr term2) with | LetIn (na1,b1,t1,c'1), LetIn (na2,b2,t2,c'2) -> let f1 i = (* FO *) ise_and i @@ -605,8 +621,8 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty evar_conv_x ts (push_rel (RelDecl.LocalDef (na,b,t)) env) i pbty c'1 c'2); (fun i -> exact_ise_stack2 env i (evar_conv_x ts) sk1 sk2)] and f2 i = - let out1 = whd_betaiota_deltazeta_for_iota_state (fst ts) env i csts1 (v1,sk1) - and out2 = whd_betaiota_deltazeta_for_iota_state (fst ts) env i csts2 (v2,sk2) + let out1 = whd_betaiota_deltazeta_for_iota_state (fst ts) env i csts1 (EConstr.of_constr v1,sk1) + and out2 = whd_betaiota_deltazeta_for_iota_state (fst ts) env i csts2 (EConstr.of_constr v2,sk2) in evar_eqappr_x ts env i pbty out1 out2 in ise_try evd [f1; f2] @@ -618,8 +634,8 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty [(fun i -> evar_conv_x ts env i CONV c c'); (fun i -> exact_ise_stack2 env i (evar_conv_x ts) sk1 sk2)] and f2 i = - let out1 = whd_betaiota_deltazeta_for_iota_state (fst ts) env i csts1 (v1,sk1) - and out2 = whd_betaiota_deltazeta_for_iota_state (fst ts) env i csts2 (v2,sk2) + let out1 = whd_betaiota_deltazeta_for_iota_state (fst ts) env i csts1 (EConstr.of_constr v1,sk1) + and out2 = whd_betaiota_deltazeta_for_iota_state (fst ts) env i csts2 (EConstr.of_constr v2,sk2) in evar_eqappr_x ts env i pbty out1 out2 in ise_try evd [f1; f2] @@ -627,7 +643,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty (* Catch the p.c ~= p c' cases *) | Proj (p,c), Const (p',u) when eq_constant (Projection.constant p) p' -> let res = - try Some (destApp (Retyping.expand_projection env evd p c [])) + try Some (EConstr.destApp evd (EConstr.of_constr (Retyping.expand_projection env evd p c []))) with Retyping.RetypeError _ -> None in (match res with @@ -638,7 +654,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty | Const (p,u), Proj (p',c') when eq_constant p (Projection.constant p') -> let res = - try Some (destApp (Retyping.expand_projection env evd p' c' [])) + try Some (EConstr.destApp evd (EConstr.of_constr (Retyping.expand_projection env evd p' c' []))) with Retyping.RetypeError _ -> None in (match res with @@ -653,7 +669,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty allow this identification (first-order unification of universes). Otherwise fallback to unfolding. *) - let univs = Universes.eq_constr_universes term1 term2 in + let univs = EConstr.eq_constr_universes evd term1 term2 in match univs with | Some univs -> ise_and i [(fun i -> @@ -675,7 +691,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty if the first argument is a beta-redex (expand a constant only if necessary) or the second argument is potentially usable as a canonical projection or canonical value *) - let rec is_unnamed (hd, args) = match kind_of_term hd with + let rec is_unnamed (hd, args) = match EConstr.kind i hd with | (Var _|Construct _|Ind _|Const _|Prod _|Sort _) -> Stack.not_purely_applicative args | (CoFix _|Meta _|Rel _)-> true @@ -684,7 +700,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty | Lambda _ -> assert (match args with [] -> true | _ -> false); true | LetIn (_,b,_,c) -> is_unnamed (fst (whd_betaiota_deltazeta_for_iota_state - (fst ts) env i Cst_stack.empty (subst1 b c, args))) + (fst ts) env i Cst_stack.empty (EConstr.Vars.subst1 b c, args))) | Fix _ -> true (* Partially applied fix can be the result of a whd call *) | Proj (p, _) -> Projection.unfolded p || Stack.not_purely_applicative args | Case _ | App _| Cast _ -> assert false in @@ -692,34 +708,35 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty let applicative_stack = fst (Stack.strip_app sk2) in is_unnamed (fst (whd_betaiota_deltazeta_for_iota_state - (fst ts) env i Cst_stack.empty (v2, applicative_stack))) in + (fst ts) env i Cst_stack.empty (EConstr.of_constr v2, applicative_stack))) in let rhs_is_already_stuck = rhs_is_already_stuck || rhs_is_stuck_and_unnamed () in - if (isLambda term1 || rhs_is_already_stuck) + if (EConstr.isLambda i term1 || rhs_is_already_stuck) && (not (Stack.not_purely_applicative sk1)) then evar_eqappr_x ~rhs_is_already_stuck ts env i pbty (whd_betaiota_deltazeta_for_iota_state - (fst ts) env i (Cst_stack.add_cst term1 csts1) (v1,sk1)) + (fst ts) env i (Cst_stack.add_cst term1 csts1) (EConstr.of_constr v1,sk1)) (appr2,csts2) else evar_eqappr_x ts env i pbty (appr1,csts1) (whd_betaiota_deltazeta_for_iota_state - (fst ts) env i (Cst_stack.add_cst term2 csts2) (v2,sk2)) + (fst ts) env i (Cst_stack.add_cst term2 csts2) (EConstr.of_constr v2,sk2)) in ise_try evd [f1; f2; f3] end - | Rigid, Rigid when isLambda term1 && isLambda term2 -> - let (na1,c1,c'1) = destLambda term1 in - let (na2,c2,c'2) = destLambda term2 in + | Rigid, Rigid when EConstr.isLambda evd term1 && EConstr.isLambda evd term2 -> + let (na1,c1,c'1) = EConstr.destLambda evd term1 in + let (na2,c2,c'2) = EConstr.destLambda evd term2 in + let inj = EConstr.Unsafe.to_constr in assert app_empty; ise_and evd - [(fun i -> evar_conv_x ts env i CONV c1 c2); + [(fun i -> evar_conv_x ts env i CONV (inj c1) (inj c2)); (fun i -> - let c = nf_evar i c1 in + let c = nf_evar i (inj c1) in let na = Nameops.name_max na1 na2 in - evar_conv_x ts (push_rel (RelDecl.LocalAssum (na,c)) env) i CONV c'1 c'2)] + evar_conv_x ts (push_rel (RelDecl.LocalAssum (na,c)) env) i CONV (inj c'1) (inj c'2))] | Flexible ev1, Rigid -> flex_rigid true ev1 appr1 appr2 | Rigid, Flexible ev2 -> flex_rigid false ev2 appr2 appr1 @@ -733,7 +750,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty and f4 i = evar_eqappr_x ts env i pbty (whd_betaiota_deltazeta_for_iota_state - (fst ts) env i (Cst_stack.add_cst term1 csts1) (v1,sk1)) + (fst ts) env i (Cst_stack.add_cst term1 csts1) (EConstr.of_constr v1,sk1)) (appr2,csts2) in ise_try evd [f3; f4] @@ -747,19 +764,20 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty and f4 i = evar_eqappr_x ts env i pbty (appr1,csts1) (whd_betaiota_deltazeta_for_iota_state - (fst ts) env i (Cst_stack.add_cst term2 csts2) (v2,sk2)) + (fst ts) env i (Cst_stack.add_cst term2 csts2) (EConstr.of_constr v2,sk2)) in ise_try evd [f3; f4] (* Eta-expansion *) - | Rigid, _ when isLambda term1 && (* if ever ill-typed: *) List.is_empty sk1 -> - eta env evd true sk1 term1 sk2 term2 + | Rigid, _ when EConstr.isLambda evd term1 && (* if ever ill-typed: *) List.is_empty sk1 -> + eta env evd true sk1 (EConstr.Unsafe.to_constr term1) sk2 term2 - | _, Rigid when isLambda term2 && (* if ever ill-typed: *) List.is_empty sk2 -> - eta env evd false sk2 term2 sk1 term1 + | _, Rigid when EConstr.isLambda evd term2 && (* if ever ill-typed: *) List.is_empty sk2 -> + eta env evd false sk2 (EConstr.Unsafe.to_constr term2) sk1 term1 | Rigid, Rigid -> begin - match kind_of_term term1, kind_of_term term2 with + let inj = EConstr.Unsafe.to_constr in + match EConstr.kind evd term1, EConstr.kind evd term2 with | Sort s1, Sort s2 when app_empty -> (try @@ -774,11 +792,11 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty | Prod (n1,c1,c'1), Prod (n2,c2,c'2) when app_empty -> ise_and evd - [(fun i -> evar_conv_x ts env i CONV c1 c2); + [(fun i -> evar_conv_x ts env i CONV (inj c1) (inj c2)); (fun i -> - let c = nf_evar i c1 in + let c = nf_evar i (inj c1) in let na = Nameops.name_max n1 n2 in - evar_conv_x ts (push_rel (RelDecl.LocalAssum (na,c)) env) i pbty c'1 c'2)] + evar_conv_x ts (push_rel (RelDecl.LocalAssum (na,c)) env) i pbty (inj c'1) (inj c'2))] | Rel x1, Rel x2 -> if Int.equal x1 x2 then @@ -822,10 +840,11 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty else UnifFailure (evd,NotSameHead) | (Meta _, _) | (_, Meta _) -> + let inj = EConstr.Unsafe.to_constr in begin match ise_stack2 true env evd (evar_conv_x ts) sk1 sk2 with |_, (UnifFailure _ as x) -> x - |None, Success i' -> evar_conv_x ts env i' CONV term1 term2 - |Some (sk1',sk2'), Success i' -> evar_conv_x ts env i' CONV (Stack.zip (term1,sk1')) (Stack.zip (term2,sk2')) + |None, Success i' -> evar_conv_x ts env i' CONV (inj term1) (inj term2) + |Some (sk1',sk2'), Success i' -> evar_conv_x ts env i' CONV (inj (Stack.zip i' (term1,sk1'))) (inj (Stack.zip i' (term2,sk2'))) end | (Ind _ | Sort _ | Prod _ | CoFix _ | Fix _ | Rel _ | Var _ | Const _), _ -> @@ -905,8 +924,8 @@ and eta_constructor ts env evd sk1 ((ind, i), u) sk2 term2 = (try let l1' = Stack.tail pars sk1 in let l2' = - let term = Stack.zip (term2,sk2) in - List.map (fun p -> mkProj (Projection.make p false, term)) (Array.to_list projs) + let term = Stack.zip evd (term2,sk2) in + List.map (fun p -> EConstr.mkProj (Projection.make p false, term)) (Array.to_list projs) in exact_ise_stack2 env evd (evar_conv_x (fst ts, false)) l1' (Stack.append_app_list l2' Stack.empty) @@ -938,14 +957,14 @@ let first_order_unification ts env evd (ev1,l1) (term2,l2) = let (deb2,rest2) = Array.chop (Array.length l2-Array.length l1) l2 in ise_and evd (* First compare extra args for better failure message *) - [(fun i -> ise_array2 i (fun i -> evar_conv_x ts env i CONV) rest2 l1); + [(fun i -> ise_array2 i (fun i -> evar_conv_x ts env i CONV) (Array.map EConstr.of_constr rest2) (Array.map EConstr.of_constr l1)); (fun i -> (* Then instantiate evar unless already done by unifying args *) let t2 = mkApp(term2,deb2) in if is_defined i (fst ev1) then evar_conv_x ts env i CONV t2 (mkEvar ev1) else - solve_simple_eqn ~choose:true (evar_conv_x ts) env i (None,ev1,t2))] + solve_simple_eqn ~choose:true (evar_conv_x ts) env i (None,ev1, EConstr.of_constr t2))] let choose_less_dependent_instance evk evd term args = let evi = Evd.find_undefined evd evk in @@ -1153,7 +1172,7 @@ let apply_conversion_problem_heuristic ts env evd pbty t1 t2 = let reason = ProblemBeyondCapabilities in UnifFailure (evd, CannotSolveConstraint ((pbty,env,t1,t2),reason))) | Evar (evk1,args1), Evar (evk2,args2) when Evar.equal evk1 evk2 -> - let f env evd pbty x y = is_fconv ~reds:ts pbty env evd x y in + let f env evd pbty x y = is_fconv ~reds:ts pbty env evd (EConstr.of_constr x) (EConstr.of_constr y) in Success (solve_refl ~can_drop:true f env evd (position_problem true pbty) evk1 args1 args2) | Evar ev1, Evar ev2 when app_empty -> diff --git a/pretyping/evarconv.mli b/pretyping/evarconv.mli index 14947c892..6f736e562 100644 --- a/pretyping/evarconv.mli +++ b/pretyping/evarconv.mli @@ -43,11 +43,11 @@ val check_problems_are_solved : env -> evar_map -> unit (** Check if a canonical structure is applicable *) val check_conv_record : env -> evar_map -> - constr * types Stack.t -> constr * types Stack.t -> + state -> state -> Univ.universe_context_set * (constr * constr) - * constr * constr list * (constr Stack.t * constr Stack.t) * - (constr Stack.t * types Stack.t) * - (constr Stack.t * types Stack.t) * constr * + * constr * constr list * (EConstr.t Stack.t * EConstr.t Stack.t) * + (EConstr.t Stack.t * EConstr.t Stack.t) * + (EConstr.t Stack.t * EConstr.t Stack.t) * constr * (int option * constr) (** Try to solve problems of the form ?x[args] = c by second-order diff --git a/pretyping/evardefine.ml b/pretyping/evardefine.ml index 06f619410..3982edd1c 100644 --- a/pretyping/evardefine.ml +++ b/pretyping/evardefine.ml @@ -33,7 +33,7 @@ let env_nf_evar sigma env = let env_nf_betaiotaevar sigma env = process_rel_context (fun d e -> - push_rel (RelDecl.map_constr (Reductionops.nf_betaiota sigma) d) e) env + push_rel (RelDecl.map_constr (fun c -> Reductionops.nf_betaiota sigma (EConstr.of_constr c)) d) e) env (****************************************) (* Operations on value/type constraints *) @@ -78,7 +78,7 @@ let define_pure_evar_as_product evd evk = let evi = Evd.find_undefined evd evk in let evenv = evar_env evi in let id = next_ident_away idx (ids_of_named_context (evar_context evi)) in - let concl = Reductionops.whd_all evenv evd evi.evar_concl in + let concl = Reductionops.whd_all evenv evd (EConstr.of_constr evi.evar_concl) in let s = destSort concl in let evd1,(dom,u1) = let evd = Sigma.Unsafe.of_evar_map evd in @@ -131,7 +131,7 @@ let define_pure_evar_as_lambda env evd evk = let open Context.Named.Declaration in let evi = Evd.find_undefined evd evk in let evenv = evar_env evi in - let typ = Reductionops.whd_all evenv evd (evar_concl evi) in + let typ = Reductionops.whd_all evenv evd (EConstr.of_constr (evar_concl evi)) in let evd1,(na,dom,rng) = match kind_of_term typ with | Prod (na,dom,rng) -> (evd,(na,dom,rng)) | Evar ev' -> let evd,typ = define_evar_as_product evd ev' in evd,destProd typ @@ -169,7 +169,7 @@ let define_evar_as_sort env evd (ev,args) = let evd, u = new_univ_variable univ_rigid evd in let evi = Evd.find_undefined evd ev in let s = Type u in - let concl = Reductionops.whd_all (evar_env evi) evd evi.evar_concl in + let concl = Reductionops.whd_all (evar_env evi) evd (EConstr.of_constr evi.evar_concl) in let sort = destSort concl in let evd' = Evd.define ev (mkSort s) evd in Evd.set_leq_sort env evd' (Type (Univ.super u)) sort, s @@ -181,7 +181,7 @@ let define_evar_as_sort env evd (ev,args) = let split_tycon loc env evd tycon = let rec real_split evd c = - let t = Reductionops.whd_all env evd c in + let t = Reductionops.whd_all env evd (EConstr.of_constr c) in match kind_of_term t with | Prod (na,dom,rng) -> evd, (na, dom, rng) | Evar ev (* ev is undefined because of whd_all *) -> diff --git a/pretyping/evarsolve.ml b/pretyping/evarsolve.ml index ea3ab17a7..17bb1406e 100644 --- a/pretyping/evarsolve.ml +++ b/pretyping/evarsolve.ml @@ -145,7 +145,7 @@ let recheck_applications conv_algo env evdref t = let argsty = Array.map (fun x -> aux env x; Retyping.get_type_of env !evdref x) args in let rec aux i ty = if i < Array.length argsty then - match kind_of_term (whd_all env !evdref ty) with + match kind_of_term (whd_all env !evdref (EConstr.of_constr ty)) with | Prod (na, dom, codom) -> (match conv_algo env !evdref Reduction.CUMUL argsty.(i) dom with | Success evd -> evdref := evd; @@ -525,7 +525,7 @@ let solve_pattern_eqn env sigma l c = l c in (* Warning: we may miss some opportunity to eta-reduce more since c' is not in normal form *) - shrink_eta c' + shrink_eta (EConstr.of_constr c') (*****************************************) (* Refining/solving unification problems *) @@ -683,7 +683,7 @@ let find_projectable_constructor env evd cstr k args cstr_subst = List.filter (fun (args',id) -> (* is_conv is maybe too strong (and source of useless computation) *) (* (at least expansion of aliases is needed) *) - Array.for_all2 (is_conv env evd) args args') l in + Array.for_all2 (fun c1 c2 -> is_conv env evd (EConstr.of_constr c1) (EConstr.of_constr c2)) args args') l in List.map snd l with Not_found -> [] @@ -808,7 +808,7 @@ let rec do_projection_effects define_fun env ty evd = function let evd = Evd.define evk (mkVar id) evd in (* TODO: simplify constraints involving evk *) let evd = do_projection_effects define_fun env ty evd p in - let ty = whd_all env evd (Lazy.force ty) in + let ty = whd_all env evd (EConstr.of_constr (Lazy.force ty)) in if not (isSort ty) then (* Don't try to instantiate if a sort because if evar_concl is an evar it may commit to a univ level which is not the right @@ -1484,7 +1484,7 @@ let rec invert_definition conv_algo choose env evd pbty (evk,argsv as ev) rhs = EConstr.Unsafe.to_constr (map_constr_with_full_binders !evdref (fun d (env,k) -> push_rel d env, k+1) self envk (EConstr.of_constr t)) in - let rhs = whd_beta evd rhs (* heuristic *) in + let rhs = whd_beta evd (EConstr.of_constr rhs) (* heuristic *) in let fast rhs = let filter_ctxt = evar_filtered_context evi in let names = ref Idset.empty in @@ -1566,10 +1566,10 @@ and evar_define conv_algo ?(choose=false) env evd pbty (evk,argsv as ev) rhs = raise e | OccurCheckIn (evd,rhs) -> (* last chance: rhs actually reduces to ev *) - let c = whd_all env evd rhs in + let c = whd_all env evd (EConstr.of_constr rhs) in match kind_of_term c with | Evar (evk',argsv2) when Evar.equal evk evk' -> - solve_refl (fun env sigma pb c c' -> is_fconv pb env sigma c c') + solve_refl (fun env sigma pb c c' -> is_fconv pb env sigma (EConstr.of_constr c) (EConstr.of_constr c')) env evd pbty evk argsv argsv2 | _ -> raise (OccurCheckIn (evd,rhs)) @@ -1638,5 +1638,5 @@ let solve_simple_eqn conv_algo ?(choose=false) env evd (pbty,(evk1,args1 as ev1) | IllTypedInstance (env,t,u) -> UnifFailure (evd,InstanceNotSameType (evk1,env,t,u)) | IncompatibleCandidates -> - UnifFailure (evd,ConversionFailed (env,mkEvar ev1,t2)) + UnifFailure (evd,ConversionFailed (env,mkEvar ev1, EConstr.Unsafe.to_constr t2)) diff --git a/pretyping/evarsolve.mli b/pretyping/evarsolve.mli index cf059febf..70e94b4dc 100644 --- a/pretyping/evarsolve.mli +++ b/pretyping/evarsolve.mli @@ -52,7 +52,7 @@ val solve_evar_evar : ?force:bool -> env -> evar_map -> bool option -> existential -> existential -> evar_map val solve_simple_eqn : conv_fun -> ?choose:bool -> env -> evar_map -> - bool option * existential * constr -> unification_result + bool option * existential * EConstr.t -> unification_result val reconsider_conv_pbs : conv_fun -> evar_map -> unification_result diff --git a/pretyping/indrec.ml b/pretyping/indrec.ml index 9cf91a947..4025ca8b8 100644 --- a/pretyping/indrec.ml +++ b/pretyping/indrec.ml @@ -153,7 +153,9 @@ let type_rec_branch is_rec dep env sigma (vargs,depPvect,decP) tyi cs recargs = let nparams = List.length vargs in let process_pos env depK pk = let rec prec env i sign p = - let p',largs = whd_allnolet_stack env sigma p in + let p',largs = whd_allnolet_stack env sigma (EConstr.of_constr p) in + let p' = EConstr.Unsafe.to_constr p' in + let largs = List.map EConstr.Unsafe.to_constr largs in match kind_of_term p' with | Prod (n,t,c) -> let d = LocalAssum (n,t) in @@ -170,7 +172,7 @@ let type_rec_branch is_rec dep env sigma (vargs,depPvect,decP) tyi cs recargs = else base | _ -> - let t' = whd_all env sigma p in + let t' = whd_all env sigma (EConstr.of_constr p) in if Term.eq_constr p' t' then assert false else prec env i sign t' in @@ -229,7 +231,9 @@ let type_rec_branch is_rec dep env sigma (vargs,depPvect,decP) tyi cs recargs = let make_rec_branch_arg env sigma (nparrec,fvect,decF) f cstr recargs = let process_pos env fk = let rec prec env i hyps p = - let p',largs = whd_allnolet_stack env sigma p in + let p',largs = whd_allnolet_stack env sigma (EConstr.of_constr p) in + let p' = EConstr.Unsafe.to_constr p' in + let largs = List.map EConstr.Unsafe.to_constr largs in match kind_of_term p' with | Prod (n,t,c) -> let d = LocalAssum (n,t) in @@ -242,7 +246,7 @@ let make_rec_branch_arg env sigma (nparrec,fvect,decF) f cstr recargs = and arg = appvect (mkRel (i+1), Context.Rel.to_extended_vect 0 hyps) in applist(lift i fk,realargs@[arg]) | _ -> - let t' = whd_all env sigma p in + let t' = whd_all env sigma (EConstr.of_constr p) in if Term.eq_constr t' p' then assert false else prec env i hyps t' in @@ -261,7 +265,7 @@ let make_rec_branch_arg env sigma (nparrec,fvect,decF) f cstr recargs = | None -> mkLambda_name env (n,t,process_constr (push_rel d env) (i+1) - (whd_beta Evd.empty (applist (lift 1 f, [(mkRel 1)]))) + (whd_beta Evd.empty (EConstr.of_constr (applist (lift 1 f, [(mkRel 1)])))) (cprest,rest)) | Some(_,f_0) -> let nF = lift (i+1+decF) f_0 in @@ -269,7 +273,7 @@ let make_rec_branch_arg env sigma (nparrec,fvect,decF) f cstr recargs = let arg = process_pos env' nF (lift 1 t) in mkLambda_name env (n,t,process_constr env' (i+1) - (whd_beta Evd.empty (applist (lift 1 f, [(mkRel 1); arg]))) + (whd_beta Evd.empty (EConstr.of_constr (applist (lift 1 f, [(mkRel 1); arg])))) (cprest,rest))) | (LocalDef (n,c,t) as d)::cprest, rest -> mkLetIn diff --git a/pretyping/inductiveops.ml b/pretyping/inductiveops.ml index 29f57144a..a3cca2ad8 100644 --- a/pretyping/inductiveops.ml +++ b/pretyping/inductiveops.ml @@ -444,15 +444,17 @@ let build_branch_type env dep p cs = (**************************************************) -let extract_mrectype t = - let (t, l) = decompose_app t in - match kind_of_term t with - | Ind ind -> (ind, l) +let extract_mrectype sigma t = + let open EConstr in + let (t, l) = decompose_app sigma t in + match EConstr.kind sigma t with + | Ind ind -> (ind, List.map EConstr.Unsafe.to_constr l) | _ -> raise Not_found let find_mrectype_vect env sigma c = - let (t, l) = decompose_appvect (whd_all env sigma c) in - match kind_of_term t with + let open EConstr in + let (t, l) = Termops.decompose_app_vect sigma (EConstr.of_constr (whd_all env sigma c)) in + match EConstr.kind sigma (EConstr.of_constr t) with | Ind ind -> (ind, l) | _ -> raise Not_found @@ -460,28 +462,34 @@ let find_mrectype env sigma c = let (ind, v) = find_mrectype_vect env sigma c in (ind, Array.to_list v) let find_rectype env sigma c = - let (t, l) = decompose_app (whd_all env sigma c) in - match kind_of_term t with + let open EConstr in + let (t, l) = decompose_app sigma (EConstr.of_constr (whd_all env sigma c)) in + match EConstr.kind sigma t with | Ind (ind,u as indu) -> let (mib,mip) = Inductive.lookup_mind_specif env ind in if mib.mind_nparams > List.length l then raise Not_found; + let l = List.map EConstr.Unsafe.to_constr l in let (par,rargs) = List.chop mib.mind_nparams l in IndType((indu, par),rargs) | _ -> raise Not_found let find_inductive env sigma c = - let (t, l) = decompose_app (whd_all env sigma c) in - match kind_of_term t with + let open EConstr in + let (t, l) = decompose_app sigma (EConstr.of_constr (whd_all env sigma c)) in + match EConstr.kind sigma t with | Ind ind when (fst (Inductive.lookup_mind_specif env (fst ind))).mind_finite <> Decl_kinds.CoFinite -> + let l = List.map EConstr.Unsafe.to_constr l in (ind, l) | _ -> raise Not_found let find_coinductive env sigma c = - let (t, l) = decompose_app (whd_all env sigma c) in - match kind_of_term t with + let open EConstr in + let (t, l) = decompose_app sigma (EConstr.of_constr (whd_all env sigma c)) in + match EConstr.kind sigma t with | Ind ind when (fst (Inductive.lookup_mind_specif env (fst ind))).mind_finite == Decl_kinds.CoFinite -> + let l = List.map EConstr.Unsafe.to_constr l in (ind, l) | _ -> raise Not_found @@ -490,12 +498,12 @@ let find_coinductive env sigma c = (* find appropriate names for pattern variables. Useful in the Case and Inversion (case_then_using et case_nodep_then_using) tactics. *) -let is_predicate_explicitly_dep env pred arsign = +let is_predicate_explicitly_dep env sigma pred arsign = let rec srec env pval arsign = - let pv' = whd_all env Evd.empty pval in - match kind_of_term pv', arsign with + let pv' = EConstr.of_constr (whd_all env sigma pval) in + match EConstr.kind sigma pv', arsign with | Lambda (na,t,b), (LocalAssum _)::arsign -> - srec (push_rel_assum (na,t) env) b arsign + srec (push_rel_assum (na, EConstr.Unsafe.to_constr t) env) b arsign | Lambda (na,_,t), _ -> (* The following code has an impact on the introduction names @@ -525,11 +533,11 @@ let is_predicate_explicitly_dep env pred arsign = | _ -> anomaly (Pp.str "Non eta-expanded dep-expanded \"match\" predicate") in - srec env pred arsign + srec env (EConstr.of_constr pred) arsign -let is_elim_predicate_explicitly_dependent env pred indf = +let is_elim_predicate_explicitly_dependent env sigma pred indf = let arsign,_ = get_arity env indf in - is_predicate_explicitly_dep env pred arsign + is_predicate_explicitly_dep env sigma pred arsign let set_names env n brty = let (ctxt,cl) = decompose_prod_n_assum n brty in @@ -545,7 +553,7 @@ let set_pattern_names env ind brv = mip.mind_nf_lc in Array.map2 (set_names env) arities brv -let type_case_branches_with_names env indspec p c = +let type_case_branches_with_names env sigma indspec p c = let (ind,args) = indspec in let (mib,mip as specif) = Inductive.lookup_mind_specif env (fst ind) in let nparams = mib.mind_nparams in @@ -554,7 +562,7 @@ let type_case_branches_with_names env indspec p c = (* Build case type *) let conclty = lambda_appvect_assum (mip.mind_nrealdecls+1) p (Array.of_list (realargs@[c])) in (* Adjust names *) - if is_elim_predicate_explicitly_dependent env p (ind,params) then + if is_elim_predicate_explicitly_dependent env sigma p (ind,params) then (set_pattern_names env (fst ind) lbrty, conclty) else (lbrty, conclty) @@ -600,7 +608,7 @@ let type_of_inductive_knowing_conclusion env sigma ((mib,mip),u) conclty = match mip.mind_arity with | RegularArity s -> sigma, subst_instance_constr u s.mind_user_arity | TemplateArity ar -> - let _,scl = Reduction.dest_arity env conclty in + let _,scl = splay_arity env sigma conclty in let ctx = List.rev mip.mind_arity_ctxt in let evdref = ref sigma in let ctx = @@ -609,6 +617,7 @@ let type_of_inductive_knowing_conclusion env sigma ((mib,mip),u) conclty = !evdref, mkArity (List.rev ctx,scl) let type_of_projection_knowing_arg env sigma p c ty = + let c = EConstr.Unsafe.to_constr c in let IndType(pars,realargs) = try find_rectype env sigma ty with Not_found -> diff --git a/pretyping/inductiveops.mli b/pretyping/inductiveops.mli index 7bd616591..1cfdef6e5 100644 --- a/pretyping/inductiveops.mli +++ b/pretyping/inductiveops.mli @@ -130,7 +130,7 @@ val has_dependent_elim : mutual_inductive_body -> bool val projection_nparams : projection -> int val projection_nparams_env : env -> projection -> int val type_of_projection_knowing_arg : env -> evar_map -> Projection.t -> - constr -> types -> types + EConstr.t -> EConstr.types -> types (** Extract information from an inductive family *) @@ -161,12 +161,12 @@ val make_arity : env -> bool -> inductive_family -> sorts -> types val build_branch_type : env -> bool -> constr -> constructor_summary -> types (** Raise [Not_found] if not given a valid inductive type *) -val extract_mrectype : constr -> pinductive * constr list -val find_mrectype : env -> evar_map -> types -> pinductive * constr list -val find_mrectype_vect : env -> evar_map -> types -> pinductive * constr array -val find_rectype : env -> evar_map -> types -> inductive_type -val find_inductive : env -> evar_map -> types -> pinductive * constr list -val find_coinductive : env -> evar_map -> types -> pinductive * constr list +val extract_mrectype : evar_map -> EConstr.t -> pinductive * constr list +val find_mrectype : env -> evar_map -> EConstr.types -> pinductive * constr list +val find_mrectype_vect : env -> evar_map -> EConstr.types -> pinductive * constr array +val find_rectype : env -> evar_map -> EConstr.types -> inductive_type +val find_inductive : env -> evar_map -> EConstr.types -> pinductive * constr list +val find_coinductive : env -> evar_map -> EConstr.types -> pinductive * constr list (********************) @@ -175,7 +175,7 @@ val arity_of_case_predicate : env -> inductive_family -> bool -> sorts -> types val type_case_branches_with_names : - env -> pinductive * constr list -> constr -> constr -> types array * types + env -> evar_map -> pinductive * constr list -> constr -> constr -> types array * types (** Annotation for cases *) val make_case_info : env -> inductive -> case_style -> case_info @@ -195,7 +195,7 @@ i*) (********************) val type_of_inductive_knowing_conclusion : - env -> evar_map -> Inductive.mind_specif puniverses -> types -> evar_map * types + env -> evar_map -> Inductive.mind_specif puniverses -> EConstr.types -> evar_map * types (********************) val control_only_guard : env -> types -> unit diff --git a/pretyping/nativenorm.ml b/pretyping/nativenorm.ml index 1e5f12b20..e45c920a3 100644 --- a/pretyping/nativenorm.ml +++ b/pretyping/nativenorm.ml @@ -92,13 +92,13 @@ let construct_of_constr_const env tag typ = let construct_of_constr_block = construct_of_constr false -let build_branches_type env (mind,_ as _ind) mib mip u params dep p = +let build_branches_type env sigma (mind,_ as _ind) mib mip u params dep p = let rtbl = mip.mind_reloc_tbl in (* [build_one_branch i cty] construit le type de la ieme branche (commence a 0) et les lambda correspondant aux realargs *) let build_one_branch i cty = let typi = type_constructor mind mib u cty params in - let decl,indapp = Reductionops.splay_prod env Evd.empty typi in + let decl,indapp = Reductionops.splay_prod env sigma (EConstr.of_constr typi) in let decl_with_letin,_ = decompose_prod_assum typi in let ind,cargs = find_rectype_a env indapp in let nparams = Array.length params in @@ -172,9 +172,9 @@ let branch_of_switch lvl ans bs = bs ci in Array.init (Array.length tbl) branch -let rec nf_val env v typ = +let rec nf_val env sigma v typ = match kind_of_value v with - | Vaccu accu -> nf_accu env accu + | Vaccu accu -> nf_accu env sigma accu | Vfun f -> let lvl = nb_rel env in let name,dom,codom = @@ -184,44 +184,44 @@ let rec nf_val env v typ = (Pp.strbrk "Returned a functional value in a type not recognized as a product type.") in let env = push_rel (LocalAssum (name,dom)) env in - let body = nf_val env (f (mk_rel_accu lvl)) codom in + let body = nf_val env sigma (f (mk_rel_accu lvl)) codom in mkLambda(name,dom,body) | Vconst n -> construct_of_constr_const env n typ | Vblock b -> let capp,ctyp = construct_of_constr_block env (block_tag b) typ in - let args = nf_bargs env b ctyp in + let args = nf_bargs env sigma b ctyp in mkApp(capp,args) -and nf_type env v = +and nf_type env sigma v = match kind_of_value v with - | Vaccu accu -> nf_accu env accu + | Vaccu accu -> nf_accu env sigma accu | _ -> assert false -and nf_type_sort env v = +and nf_type_sort env sigma v = match kind_of_value v with | Vaccu accu -> - let t,s = nf_accu_type env accu in + let t,s = nf_accu_type env sigma accu in let s = try destSort s with DestKO -> assert false in t, s | _ -> assert false -and nf_accu env accu = +and nf_accu env sigma accu = let atom = atom_of_accu accu in - if Int.equal (accu_nargs accu) 0 then nf_atom env atom + if Int.equal (accu_nargs accu) 0 then nf_atom env sigma atom else - let a,typ = nf_atom_type env atom in - let _, args = nf_args env accu typ in + let a,typ = nf_atom_type env sigma atom in + let _, args = nf_args env sigma accu typ in mkApp(a,Array.of_list args) -and nf_accu_type env accu = +and nf_accu_type env sigma accu = let atom = atom_of_accu accu in - if Int.equal (accu_nargs accu) 0 then nf_atom_type env atom + if Int.equal (accu_nargs accu) 0 then nf_atom_type env sigma atom else - let a,typ = nf_atom_type env atom in - let t, args = nf_args env accu typ in + let a,typ = nf_atom_type env sigma atom in + let t, args = nf_args env sigma accu typ in mkApp(a,Array.of_list args), t -and nf_args env accu t = +and nf_args env sigma accu t = let aux arg (t,l) = let _,dom,codom = try decompose_prod env t with @@ -229,13 +229,13 @@ and nf_args env accu t = CErrors.anomaly (Pp.strbrk "Returned a functional value in a type not recognized as a product type.") in - let c = nf_val env arg dom in + let c = nf_val env sigma arg dom in (subst1 c codom, c::l) in let t,l = List.fold_right aux (args_of_accu accu) (t,[]) in t, List.rev l -and nf_bargs env b t = +and nf_bargs env sigma b t = let t = ref t in let len = block_size b in Array.init len @@ -246,10 +246,10 @@ and nf_bargs env b t = CErrors.anomaly (Pp.strbrk "Returned a functional value in a type not recognized as a product type.") in - let c = nf_val env (block_field b i) dom in + let c = nf_val env sigma (block_field b i) dom in t := subst1 c codom; c) -and nf_atom env atom = +and nf_atom env sigma atom = match atom with | Arel i -> mkRel (nb_rel env - i) | Aconstant cst -> mkConstU cst @@ -257,19 +257,19 @@ and nf_atom env atom = | Asort s -> mkSort s | Avar id -> mkVar id | Aprod(n,dom,codom) -> - let dom = nf_type env dom in + let dom = nf_type env sigma dom in let vn = mk_rel_accu (nb_rel env) in let env = push_rel (LocalAssum (n,dom)) env in - let codom = nf_type env (codom vn) in + let codom = nf_type env sigma (codom vn) in mkProd(n,dom,codom) | Ameta (mv,_) -> mkMeta mv | Aevar (ev,_) -> mkEvar ev | Aproj(p,c) -> - let c = nf_accu env c in + let c = nf_accu env sigma c in mkProj(Projection.make p true,c) - | _ -> fst (nf_atom_type env atom) + | _ -> fst (nf_atom_type env sigma atom) -and nf_atom_type env atom = +and nf_atom_type env sigma atom = match atom with | Arel i -> let n = (nb_rel env - i) in @@ -283,7 +283,7 @@ and nf_atom_type env atom = | Avar id -> mkVar id, type_of_var env id | Acase(ans,accu,p,bs) -> - let a,ta = nf_accu_type env accu in + let a,ta = nf_accu_type env sigma accu in let ((mind,_),u as ind),allargs = find_rectype_a env ta in let (mib,mip) = Inductive.lookup_mind_specif env (fst ind) in let nparams = mib.mind_nparams in @@ -292,14 +292,14 @@ and nf_atom_type env atom = hnf_prod_applist env (Inductiveops.type_of_inductive env ind) (Array.to_list params) in let pT = whd_all env pT in - let dep, p = nf_predicate env ind mip params p pT in + let dep, p = nf_predicate env sigma ind mip params p pT in (* Calcul du type des branches *) - let btypes = build_branches_type env (fst ind) mib mip u params dep p in + let btypes = build_branches_type env sigma (fst ind) mib mip u params dep p in (* calcul des branches *) let bsw = branch_of_switch (nb_rel env) ans bs in let mkbranch i v = let decl,decl_with_letin,codom = btypes.(i) in - let b = nf_val (Termops.push_rels_assum decl env) v codom in + let b = nf_val (Termops.push_rels_assum decl env) sigma v codom in Termops.it_mkLambda_or_LetIn_from_no_LetIn b decl_with_letin in let branchs = Array.mapi mkbranch bsw in @@ -307,7 +307,7 @@ and nf_atom_type env atom = let ci = ans.asw_ci in mkCase(ci, p, a, branchs), tcase | Afix(tt,ft,rp,s) -> - let tt = Array.map (fun t -> nf_type env t) tt in + let tt = Array.map (fun t -> nf_type env sigma t) tt in let name = Array.map (fun _ -> (Name (id_of_string "Ffix"))) tt in let lvl = nb_rel env in let nbfix = Array.length ft in @@ -316,37 +316,37 @@ and nf_atom_type env atom = let env = push_rec_types (name,tt,[||]) env in (* We lift here because the types of arguments (in tt) will be evaluated in an environment where the fixpoints have been pushed *) - let norm_body i v = nf_val env (napply v fargs) (lift nbfix tt.(i)) in + let norm_body i v = nf_val env sigma (napply v fargs) (lift nbfix tt.(i)) in let ft = Array.mapi norm_body ft in mkFix((rp,s),(name,tt,ft)), tt.(s) | Acofix(tt,ft,s,_) | Acofixe(tt,ft,s,_) -> - let tt = Array.map (nf_type env) tt in + let tt = Array.map (nf_type env sigma) tt in let name = Array.map (fun _ -> (Name (id_of_string "Fcofix"))) tt in let lvl = nb_rel env in let fargs = mk_rels_accu lvl (Array.length ft) in let env = push_rec_types (name,tt,[||]) env in - let ft = Array.mapi (fun i v -> nf_val env (napply v fargs) tt.(i)) ft in + let ft = Array.mapi (fun i v -> nf_val env sigma (napply v fargs) tt.(i)) ft in mkCoFix(s,(name,tt,ft)), tt.(s) | Aprod(n,dom,codom) -> - let dom,s1 = nf_type_sort env dom in + let dom,s1 = nf_type_sort env sigma dom in let vn = mk_rel_accu (nb_rel env) in let env = push_rel (LocalAssum (n,dom)) env in - let codom,s2 = nf_type_sort env (codom vn) in + let codom,s2 = nf_type_sort env sigma (codom vn) in mkProd(n,dom,codom), mkSort (sort_of_product env s1 s2) | Aevar(ev,ty) -> - let ty = nf_type env ty in + let ty = nf_type env sigma ty in mkEvar ev, ty | Ameta(mv,ty) -> - let ty = nf_type env ty in + let ty = nf_type env sigma ty in mkMeta mv, ty | Aproj(p,c) -> - let c,tc = nf_accu_type env c in + let c,tc = nf_accu_type env sigma c in let cj = make_judge c tc in let uj = Typeops.judge_of_projection env (Projection.make p true) cj in uj.uj_val, uj.uj_type -and nf_predicate env ind mip params v pT = +and nf_predicate env sigma ind mip params v pT = match kind_of_value v, kind_of_term pT with | Vfun f, Prod _ -> let k = nb_rel env in @@ -358,7 +358,7 @@ and nf_predicate env ind mip params v pT = (Pp.strbrk "Returned a functional value in a type not recognized as a product type.") in let dep,body = - nf_predicate (push_rel (LocalAssum (name,dom)) env) ind mip params vb codom in + nf_predicate (push_rel (LocalAssum (name,dom)) env) sigma ind mip params vb codom in dep, mkLambda(name,dom,body) | Vfun f, _ -> let k = nb_rel env in @@ -368,9 +368,9 @@ and nf_predicate env ind mip params v pT = let rargs = Array.init n (fun i -> mkRel (n-i)) in let params = if Int.equal n 0 then params else Array.map (lift n) params in let dom = mkApp(mkIndU ind,Array.append params rargs) in - let body = nf_type (push_rel (LocalAssum (name,dom)) env) vb in + let body = nf_type (push_rel (LocalAssum (name,dom)) env) sigma vb in true, mkLambda(name,dom,body) - | _, _ -> false, nf_type env v + | _, _ -> false, nf_type env sigma v let evars_of_evar_map sigma = { Nativelambda.evars_val = Evd.existential_opt_value sigma; @@ -382,13 +382,12 @@ let native_norm env sigma c ty = error "Native_compute reduction has been disabled at configure time." else let penv = Environ.pre_env env in - let sigma = evars_of_evar_map sigma in (* Format.eprintf "Numbers of free variables (named): %i\n" (List.length vl1); Format.eprintf "Numbers of free variables (rel): %i\n" (List.length vl2); *) let ml_filename, prefix = Nativelib.get_ml_filename () in - let code, upd = mk_norm_code penv sigma prefix c in + let code, upd = mk_norm_code penv (evars_of_evar_map sigma) prefix c in match Nativelib.compile ml_filename code with | true, fn -> if !Flags.debug then Feedback.msg_debug (Pp.str "Running norm ..."); @@ -397,7 +396,7 @@ let native_norm env sigma c ty = let t1 = Sys.time () in let time_info = Format.sprintf "Evaluation done in %.5f@." (t1 -. t0) in if !Flags.debug then Feedback.msg_debug (Pp.str time_info); - let res = nf_val env !Nativelib.rt1 ty in + let res = nf_val env sigma !Nativelib.rt1 ty in let t2 = Sys.time () in let time_info = Format.sprintf "Reification done in %.5f@." (t2 -. t1) in if !Flags.debug then Feedback.msg_debug (Pp.str time_info); diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml index 9b572f376..3a6d4f36c 100644 --- a/pretyping/pretyping.ml +++ b/pretyping/pretyping.ml @@ -700,7 +700,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre if Int.equal npars 0 then [] else try - let IndType (indf, args) = find_rectype env.ExtraEnv.env !evdref ty in + let IndType (indf, args) = find_rectype env.ExtraEnv.env !evdref (EConstr.of_constr ty) in let ((ind',u'),pars) = dest_ind_family indf in if eq_ind ind ind' then pars else (* Let the usual code throw an error *) [] @@ -723,7 +723,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre | c::rest -> let argloc = loc_of_glob_constr c in let resj = evd_comb1 (Coercion.inh_app_fun resolve_tc env.ExtraEnv.env) evdref resj in - let resty = whd_all env.ExtraEnv.env !evdref resj.uj_type in + let resty = whd_all env.ExtraEnv.env !evdref (EConstr.of_constr resj.uj_type) in match kind_of_term resty with | Prod (na,c1,c2) -> let tycon = Some c1 in @@ -834,7 +834,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre | GLetTuple (loc,nal,(na,po),c,d) -> let cj = pretype empty_tycon env evdref lvar c in let (IndType (indf,realargs)) = - try find_rectype env.ExtraEnv.env !evdref cj.uj_type + try find_rectype env.ExtraEnv.env !evdref (EConstr.of_constr cj.uj_type) with Not_found -> let cloc = loc_of_glob_constr c in error_case_not_inductive ~loc:cloc env.ExtraEnv.env !evdref cj @@ -894,7 +894,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre (Array.to_list cs.cs_concl_realargs) @[build_dependent_constructor cs] in let lp = lift cs.cs_nargs p in - let fty = hnf_lam_applist env.ExtraEnv.env !evdref lp inst in + let fty = hnf_lam_applist env.ExtraEnv.env !evdref (EConstr.of_constr lp) (List.map EConstr.of_constr inst) in let fj = pretype (mk_tycon fty) env_f evdref lvar d in let v = let ind,_ = dest_ind_family indf in @@ -924,7 +924,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre | GIf (loc,c,(na,po),b1,b2) -> let cj = pretype empty_tycon env evdref lvar c in let (IndType (indf,realargs)) = - try find_rectype env.ExtraEnv.env !evdref cj.uj_type + try find_rectype env.ExtraEnv.env !evdref (EConstr.of_constr cj.uj_type) with Not_found -> let cloc = loc_of_glob_constr c in error_case_not_inductive ~loc:cloc env.ExtraEnv.env !evdref cj in @@ -948,7 +948,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre let pj = pretype_type empty_valcon env_p evdref lvar p in let ccl = nf_evar !evdref pj.utj_val in let pred = it_mkLambda_or_LetIn ccl psign in - let typ = lift (- nar) (beta_applist (pred,[cj.uj_val])) in + let typ = lift (- nar) (beta_applist !evdref (EConstr.of_constr pred,[EConstr.of_constr cj.uj_val])) in pred, typ | None -> let p = match tycon with @@ -963,7 +963,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre let f cs b = let n = Context.Rel.length cs.cs_args in let pi = lift n pred in (* liftn n 2 pred ? *) - let pi = beta_applist (pi, [build_dependent_constructor cs]) in + let pi = beta_applist !evdref (EConstr.of_constr pi, [EConstr.of_constr (build_dependent_constructor cs)]) in let csgn = if not !allow_anonymous_refs then List.map (set_name Anonymous) cs.cs_args @@ -1046,11 +1046,11 @@ and pretype_instance k0 resolve_tc env evdref lvar loc hyps evk update = with Not_found -> try let (n,_,t') = lookup_rel_id id (rel_context env) in - if is_conv env.ExtraEnv.env !evdref t t' then mkRel n, update else raise Not_found + if is_conv env.ExtraEnv.env !evdref (EConstr.of_constr t) (EConstr.of_constr t') then mkRel n, update else raise Not_found with Not_found -> try let t' = env |> lookup_named id |> NamedDecl.get_type in - if is_conv env.ExtraEnv.env !evdref t t' then mkVar id, update else raise Not_found + if is_conv env.ExtraEnv.env !evdref (EConstr.of_constr t) (EConstr.of_constr t') then mkVar id, update else raise Not_found with Not_found -> user_err ~loc (str "Cannot interpret " ++ pr_existential_key !evdref evk ++ @@ -1068,7 +1068,7 @@ and pretype_type k0 resolve_tc valcon (env : ExtraEnv.t) evdref lvar = function let s = let sigma = !evdref in let t = Retyping.get_type_of env.ExtraEnv.env sigma v in - match kind_of_term (whd_all env.ExtraEnv.env sigma t) with + match kind_of_term (whd_all env.ExtraEnv.env sigma (EConstr.of_constr t)) with | Sort s -> s | Evar ev when is_Type (existential_type sigma ev) -> evd_comb1 (define_evar_as_sort env.ExtraEnv.env) evdref ev diff --git a/pretyping/recordops.ml b/pretyping/recordops.ml index e897d5f5c..062e4a068 100644 --- a/pretyping/recordops.ml +++ b/pretyping/recordops.ml @@ -202,7 +202,7 @@ let compute_canonical_projections warn (con,ind) = let v = (mkConstU (con,u)) in let ctx = Univ.ContextSet.of_context ctx in let c = Environ.constant_value_in env (con,u) in - let lt,t = Reductionops.splay_lam env Evd.empty c in + let lt,t = Reductionops.splay_lam env Evd.empty (EConstr.of_constr c) in let lt = List.rev_map snd lt in let args = snd (decompose_app t) in let { s_EXPECTEDPARAM = p; s_PROJ = lpj; s_PROJKIND = kl } = @@ -302,7 +302,7 @@ let check_and_decompose_canonical_structure ref = let vc = match Environ.constant_opt_value_in env (sp, u) with | Some vc -> vc | None -> error_not_structure ref in - let body = snd (splay_lam (Global.env()) Evd.empty vc) in + let body = snd (splay_lam (Global.env()) Evd.empty (EConstr.of_constr vc)) (** FIXME *) in let f,args = match kind_of_term body with | App (f,args) -> f,args | _ -> error_not_structure ref in @@ -323,6 +323,7 @@ let lookup_canonical_conversion (proj,pat) = let is_open_canonical_projection env sigma (c,args) = try + let c = EConstr.Unsafe.to_constr c in let ref = global_of_constr c in let n = find_projection_nparams ref in (** Check if there is some canonical projection attached to this structure *) diff --git a/pretyping/recordops.mli b/pretyping/recordops.mli index 4a176760c..405963a9c 100644 --- a/pretyping/recordops.mli +++ b/pretyping/recordops.mli @@ -72,6 +72,6 @@ val pr_cs_pattern : cs_pattern -> Pp.std_ppcmds val lookup_canonical_conversion : (global_reference * cs_pattern) -> constr * obj_typ val declare_canonical_structure : global_reference -> unit val is_open_canonical_projection : - Environ.env -> Evd.evar_map -> (constr * constr Reductionops.Stack.t) -> bool + Environ.env -> Evd.evar_map -> Reductionops.state -> bool val canonical_projections : unit -> ((global_reference * cs_pattern) * obj_typ) list diff --git a/pretyping/reductionops.ml b/pretyping/reductionops.ml index 820974888..69d47e8e6 100644 --- a/pretyping/reductionops.ml +++ b/pretyping/reductionops.ml @@ -156,6 +156,7 @@ end (** Machinery about stack of unfolded constants *) module Cst_stack = struct + open EConstr (** constant * params * args - constant applied to params = term in head applied to args @@ -191,8 +192,8 @@ module Cst_stack = struct | (cst,params,[])::_ -> Some(cst,params) | _ -> None - let reference t = match best_cst t with - | Some (c, _) when Term.isConst c -> Some (fst (Term.destConst c)) + let reference sigma t = match best_cst t with + | Some (c, _) when isConst sigma c -> Some (fst (destConst sigma c)) | _ -> None (** [best_replace d cst_l c] makes the best replacement for [d] @@ -201,14 +202,14 @@ module Cst_stack = struct let reconstruct_head = List.fold_left (fun t (i,args) -> mkApp (t,Array.sub args i (Array.length args - i))) in List.fold_right - (fun (cst,params,args) t -> Termops.replace_term sigma - (EConstr.of_constr (reconstruct_head d args)) - (EConstr.of_constr (applist (cst, List.rev params))) - (EConstr.of_constr t)) cst_l c + (fun (cst,params,args) t -> EConstr.of_constr (Termops.replace_term sigma + (reconstruct_head d args) + (applist (cst, List.rev params)) + t)) cst_l c let pr l = let open Pp in - let p_c = Termops.print_constr in + let p_c c = Termops.print_constr (EConstr.Unsafe.to_constr c) in prlist_with_sep pr_semicolon (fun (c,params,args) -> hov 1 (str"(" ++ p_c c ++ str ")" ++ spc () ++ pr_sequence p_c params ++ spc () ++ str "(args:" ++ @@ -220,6 +221,7 @@ end (** The type of (machine) stacks (= lambda-bar-calculus' contexts) *) module Stack : sig + open EConstr type 'a app_node val pr_app_node : ('a -> Pp.std_ppcmds) -> 'a app_node -> Pp.std_ppcmds @@ -231,7 +233,7 @@ sig | App of 'a app_node | Case of case_info * 'a * 'a array * Cst_stack.t | Proj of int * int * projection * Cst_stack.t - | Fix of fixpoint * 'a t * Cst_stack.t + | Fix of ('a, 'a) pfixpoint * 'a t * Cst_stack.t | Cst of cst_member * int * int list * 'a t * Cst_stack.t | Shift of int | Update of 'a @@ -242,10 +244,10 @@ sig val append_app : 'a array -> 'a t -> 'a t val decomp : 'a t -> ('a * 'a t) option val decomp_node_last : 'a app_node -> 'a t -> ('a * 'a t) - val equal : ('a * int -> 'a * int -> bool) -> (fixpoint * int -> fixpoint * int -> bool) + val equal : ('a * int -> 'a * int -> bool) -> (('a, 'a) pfixpoint * int -> ('a, 'a) pfixpoint * int -> bool) -> 'a t -> 'a t -> (int * int) option val compare_shape : 'a t -> 'a t -> bool - val map : (constr -> constr) -> constr t -> constr t + val map : ('a -> 'a) -> 'a t -> 'a t val fold2 : ('a -> constr -> constr -> 'a) -> 'a -> constr t -> constr t -> 'a * int * int val append_app_list : 'a list -> 'a t -> 'a t @@ -258,10 +260,11 @@ sig val args_size : 'a t -> int val tail : int -> 'a t -> 'a t val nth : 'a t -> int -> 'a - val best_state : constr * constr t -> Cst_stack.t -> constr * constr t - val zip : ?refold:bool -> constr * constr t -> constr + val best_state : evar_map -> constr * constr t -> Cst_stack.t -> constr * constr t + val zip : ?refold:bool -> evar_map -> constr * constr t -> constr end = struct + open EConstr type 'a app_node = int * 'a array * int (* first releavnt position, arguments, last relevant position *) @@ -286,7 +289,7 @@ struct | App of 'a app_node | Case of Term.case_info * 'a * 'a array * Cst_stack.t | Proj of int * int * projection * Cst_stack.t - | Fix of fixpoint * 'a t * Cst_stack.t + | Fix of ('a, 'a) pfixpoint * 'a t * Cst_stack.t | Cst of cst_member * int * int list * 'a t * Cst_stack.t | Shift of int | Update of 'a @@ -305,7 +308,7 @@ struct str "ZProj(" ++ int n ++ pr_comma () ++ int m ++ pr_comma () ++ pr_con (Projection.constant p) ++ str ")" | Fix (f,args,cst) -> - str "ZFix(" ++ Termops.pr_fix Termops.print_constr f + str "ZFix(" ++ Termops.pr_fix pr_c f ++ pr_comma () ++ pr pr_c args ++ str ")" | Cst (mem,curr,remains,params,cst_l) -> str "ZCst(" ++ pr_cst_member pr_c mem ++ pr_comma () ++ int curr @@ -533,11 +536,11 @@ struct | None -> raise Not_found (** This function breaks the abstraction of Cst_stack ! *) - let best_state (_,sk as s) l = + let best_state sigma (_,sk as s) l = let rec aux sk def = function |(cst, params, []) -> (cst, append_app_list (List.rev params) sk) |(cst, params, (i,t)::q) -> match decomp sk with - | Some (el,sk') when Constr.equal el t.(i) -> + | Some (el,sk') when EConstr.eq_constr sigma el t.(i) -> if i = pred (Array.length t) then aux sk' def (cst, params, q) else aux sk' def (cst, params, (succ i,t)::q) @@ -552,53 +555,66 @@ struct | Some (hd, sk) -> mkProj (p, hd), sk | None -> assert false - let rec zip ?(refold=false) = function + let zip ?(refold=false) sigma s = + let rec zip = function | f, [] -> f | f, (App (i,a,j) :: s) -> let a' = if Int.equal i 0 && Int.equal j (Array.length a - 1) then a else Array.sub a i (j - i + 1) in - zip ~refold (mkApp (f, a'), s) + zip (mkApp (f, a'), s) | f, (Case (ci,rt,br,cst_l)::s) when refold -> - zip ~refold (best_state (mkCase (ci,rt,f,br), s) cst_l) - | f, (Case (ci,rt,br,_)::s) -> zip ~refold (mkCase (ci,rt,f,br), s) + zip (best_state sigma (mkCase (ci,rt,f,br), s) cst_l) + | f, (Case (ci,rt,br,_)::s) -> zip (mkCase (ci,rt,f,br), s) | f, (Fix (fix,st,cst_l)::s) when refold -> - zip ~refold (best_state (mkFix fix, st @ (append_app [|f|] s)) cst_l) - | f, (Fix (fix,st,_)::s) -> zip ~refold + zip (best_state sigma (mkFix fix, st @ (append_app [|f|] s)) cst_l) + | f, (Fix (fix,st,_)::s) -> zip (mkFix fix, st @ (append_app [|f|] s)) | f, (Cst (cst,_,_,params,cst_l)::s) when refold -> - zip ~refold (best_state (constr_of_cst_member cst (params @ (append_app [|f|] s))) cst_l) + zip (best_state sigma (constr_of_cst_member cst (params @ (append_app [|f|] s))) cst_l) | f, (Cst (cst,_,_,params,_)::s) -> - zip ~refold (constr_of_cst_member cst (params @ (append_app [|f|] s))) - | f, (Shift n::s) -> zip ~refold (lift n f, s) + zip (constr_of_cst_member cst (params @ (append_app [|f|] s))) + | f, (Shift n::s) -> zip (Vars.lift n f, s) | f, (Proj (n,m,p,cst_l)::s) when refold -> - zip ~refold (best_state (mkProj (p,f),s) cst_l) - | f, (Proj (n,m,p,_)::s) -> zip ~refold (mkProj (p,f),s) + zip (best_state sigma (mkProj (p,f),s) cst_l) + | f, (Proj (n,m,p,_)::s) -> zip (mkProj (p,f),s) | _, (Update _::_) -> assert false + in + zip s + end (** The type of (machine) states (= lambda-bar-calculus' cuts) *) -type state = constr * constr Stack.t +type state = EConstr.t * EConstr.t Stack.t -type contextual_reduction_function = env -> evar_map -> constr -> constr -type reduction_function = contextual_reduction_function -type local_reduction_function = evar_map -> constr -> constr -type e_reduction_function = { e_redfun : 'r. env -> 'r Sigma.t -> constr -> (constr, 'r) Sigma.sigma } +type contextual_reduction_function = env -> evar_map -> EConstr.t -> constr +type reduction_function = contextual_reduction_function +type local_reduction_function = evar_map -> EConstr.t -> constr +type e_reduction_function = { e_redfun : 'r. env -> 'r Sigma.t -> EConstr.t -> (constr, 'r) Sigma.sigma } -type contextual_stack_reduction_function = - env -> evar_map -> constr -> constr * constr list -type stack_reduction_function = contextual_stack_reduction_function +type contextual_stack_reduction_function = + env -> evar_map -> EConstr.t -> EConstr.t * EConstr.t list +type stack_reduction_function = contextual_stack_reduction_function type local_stack_reduction_function = - evar_map -> constr -> constr * constr list + evar_map -> EConstr.t -> EConstr.t * EConstr.t list -type contextual_state_reduction_function = - env -> evar_map -> state -> state -type state_reduction_function = contextual_state_reduction_function +type contextual_state_reduction_function = + env -> evar_map -> state -> state +type state_reduction_function = contextual_state_reduction_function type local_state_reduction_function = evar_map -> state -> state let pr_state (tm,sk) = let open Pp in - h 0 (Termops.print_constr tm ++ str "|" ++ cut () ++ Stack.pr Termops.print_constr sk) + let pr c = Termops.print_constr (EConstr.Unsafe.to_constr c) in + h 0 (pr tm ++ str "|" ++ cut () ++ Stack.pr pr sk) + +let local_assum (na, t) = + let inj = EConstr.Unsafe.to_constr in + LocalAssum (na, inj t) + +let local_def (na, b, t) = + let inj = EConstr.Unsafe.to_constr in + LocalDef (na, inj b, inj t) (*************************************) (*** Reduction Functions Operators ***) @@ -612,19 +628,19 @@ let safe_meta_value sigma ev = let strong whdfun env sigma t = let rec strongrec env t = - let t = EConstr.of_constr (whdfun env sigma (EConstr.Unsafe.to_constr t)) in + let t = EConstr.of_constr (whdfun env sigma t) in map_constr_with_full_binders sigma push_rel strongrec env t in - EConstr.Unsafe.to_constr (strongrec env (EConstr.of_constr t)) + EConstr.Unsafe.to_constr (strongrec env t) let local_strong whdfun sigma = - let rec strongrec t = Constr.map strongrec (whdfun sigma t) in - strongrec + let rec strongrec t = EConstr.map sigma strongrec (EConstr.of_constr (whdfun sigma t)) in + fun c -> EConstr.Unsafe.to_constr (strongrec c) let rec strong_prodspine redfun sigma c = - let x = redfun sigma c in - match kind_of_term x with - | Prod (na,a,b) -> mkProd (na,a,strong_prodspine redfun sigma b) - | _ -> x + let x = EConstr.of_constr (redfun sigma c) in + match EConstr.kind sigma x with + | Prod (na,a,b) -> mkProd (na, EConstr.Unsafe.to_constr a,strong_prodspine redfun sigma b) + | _ -> EConstr.Unsafe.to_constr x (*************************************) (*** Reduction using bindingss ***) @@ -634,31 +650,36 @@ let eta = CClosure.RedFlags.mkflags [CClosure.RedFlags.fETA] (* Beta Reduction tools *) -let apply_subst recfun env refold cst_l t stack = +let apply_subst recfun env sigma refold cst_l t stack = let rec aux env cst_l t stack = - match (Stack.decomp stack,kind_of_term t) with + match (Stack.decomp stack, EConstr.kind sigma t) with | Some (h,stacktl), Lambda (_,_,c) -> let cst_l' = if refold then Cst_stack.add_param h cst_l else cst_l in aux (h::env) cst_l' c stacktl - | _ -> recfun cst_l (substl env t, stack) + | _ -> recfun sigma cst_l (EConstr.Vars.substl env t, stack) in aux env cst_l t stack -let stacklam recfun env t stack = - apply_subst (fun _ -> recfun) env false Cst_stack.empty t stack +let stacklam recfun env sigma t stack = + apply_subst (fun _ _ s -> recfun s) env sigma false Cst_stack.empty t stack + +let beta_app sigma (c,l) = + let zip s = Stack.zip sigma s in + stacklam zip [] sigma c (Stack.append_app l Stack.empty) -let beta_applist (c,l) = - stacklam Stack.zip [] c (Stack.append_app_list l Stack.empty) +let beta_applist sigma (c,l) = + let zip s = Stack.zip sigma s in + EConstr.Unsafe.to_constr (stacklam zip [] sigma c (Stack.append_app_list l Stack.empty)) (* Iota reduction tools *) type 'a miota_args = { - mP : constr; (* the result type *) - mconstr : constr; (* the constructor *) + mP : EConstr.t; (* the result type *) + mconstr : EConstr.t; (* the constructor *) mci : case_info; (* special info to re-build pattern *) mcargs : 'a list; (* the constructor's arguments *) mlf : 'a array } (* the branch code vector *) -let reducible_mind_case c = match kind_of_term c with +let reducible_mind_case sigma c = match EConstr.kind sigma c with | Construct _ | CoFix _ -> true | _ -> false @@ -672,9 +693,10 @@ let reducible_mind_case c = match kind_of_term c with f x := t. End M. Definition f := u. and say goodbye to any hope of refolding M.f this way ... *) -let magicaly_constant_of_fixbody env reference bd = function +let magicaly_constant_of_fixbody env sigma reference bd = function | Name.Anonymous -> bd | Name.Name id -> + let open EConstr in try let (cst_mod,cst_sect,_) = Constant.repr3 reference in let cst = Constant.make3 cst_mod cst_sect (Label.of_id id) in @@ -682,7 +704,7 @@ let magicaly_constant_of_fixbody env reference bd = function match constant_opt_value_in env (cst,u) with | None -> bd | Some t -> - let csts = Universes.eq_constr_universes t bd in + let csts = EConstr.eq_constr_universes sigma (EConstr.of_constr t) bd in begin match csts with | Some csts -> let subst = Universes.Constraints.fold (fun (l,d,r) acc -> @@ -696,7 +718,8 @@ let magicaly_constant_of_fixbody env reference bd = function with | Not_found -> bd -let contract_cofix ?env ?reference (bodynum,(names,types,bodies as typedbodies)) = +let contract_cofix ?env sigma ?reference (bodynum,(names,types,bodies as typedbodies)) = + let open EConstr in let nbodies = Array.length bodies in let make_Fi j = let ind = nbodies-j-1 in @@ -708,37 +731,40 @@ let contract_cofix ?env ?reference (bodynum,(names,types,bodies as typedbodies)) | Some e -> match reference with | None -> bd - | Some r -> magicaly_constant_of_fixbody e r bd names.(ind) in + | Some r -> magicaly_constant_of_fixbody e sigma r bd names.(ind) in let closure = List.init nbodies make_Fi in - substl closure bodies.(bodynum) + Vars.substl closure bodies.(bodynum) (** Similar to the "fix" case below *) let reduce_and_refold_cofix recfun env sigma refold cst_l cofix sk = + let open EConstr in let raw_answer = let env = if refold then Some env else None in - contract_cofix ?env ?reference:(Cst_stack.reference cst_l) cofix in + contract_cofix ?env sigma ?reference:(Cst_stack.reference sigma cst_l) cofix in apply_subst - (fun x (t,sk') -> + (fun sigma x (t,sk') -> let t' = if refold then Cst_stack.best_replace sigma (mkCoFix cofix) cst_l t else t in recfun x (t',sk')) - [] refold Cst_stack.empty raw_answer sk + [] sigma refold Cst_stack.empty raw_answer sk -let reduce_mind_case mia = - match kind_of_term mia.mconstr with +let reduce_mind_case sigma mia = + let open EConstr in + match EConstr.kind sigma mia.mconstr with | Construct ((ind_sp,i),u) -> (* let ncargs = (fst mia.mci).(i-1) in*) let real_cargs = List.skipn mia.mci.ci_npar mia.mcargs in applist (mia.mlf.(i-1),real_cargs) | CoFix cofix -> - let cofix_def = contract_cofix cofix in + let cofix_def = contract_cofix sigma cofix in mkCase (mia.mci, mia.mP, applist(cofix_def,mia.mcargs), mia.mlf) | _ -> assert false (* contracts fix==FIX[nl;i](A1...Ak;[F1...Fk]{B1....Bk}) to produce Bi[Fj --> FIX[nl;j](A1...Ak;[F1...Fk]{B1...Bk})] *) -let contract_fix ?env ?reference ((recindices,bodynum),(names,types,bodies as typedbodies)) = +let contract_fix ?env sigma ?reference ((recindices,bodynum),(names,types,bodies as typedbodies)) = + let open EConstr in let nbodies = Array.length recindices in let make_Fi j = let ind = nbodies-j-1 in @@ -750,26 +776,27 @@ let contract_fix ?env ?reference ((recindices,bodynum),(names,types,bodies as ty | Some e -> match reference with | None -> bd - | Some r -> magicaly_constant_of_fixbody e r bd names.(ind) in + | Some r -> magicaly_constant_of_fixbody e sigma r bd names.(ind) in let closure = List.init nbodies make_Fi in - substl closure bodies.(bodynum) + Vars.substl closure bodies.(bodynum) (** First we substitute the Rel bodynum by the fixpoint and then we try to replace the fixpoint by the best constant from [cst_l] Other rels are directly substituted by constants "magically found from the context" in contract_fix *) let reduce_and_refold_fix recfun env sigma refold cst_l fix sk = + let open EConstr in let raw_answer = let env = if refold then None else Some env in - contract_fix ?env ?reference:(Cst_stack.reference cst_l) fix in + contract_fix ?env sigma ?reference:(Cst_stack.reference sigma cst_l) fix in apply_subst - (fun x (t,sk') -> + (fun sigma x (t,sk') -> let t' = if refold then Cst_stack.best_replace sigma (mkFix fix) cst_l t else t in recfun x (t',sk')) - [] refold Cst_stack.empty raw_answer sk + [] sigma refold Cst_stack.empty raw_answer sk let fix_recarg ((recindices,bodynum),_) stack = assert (0 <= bodynum && bodynum < Array.length recindices); @@ -802,51 +829,53 @@ let _ = Goptions.declare_bool_option { Goptions.optwrite = (fun a -> debug_RAKAM:=a); } -let equal_stacks (x, l) (y, l') = - let f_equal (x,lft1) (y,lft2) = Constr.equal (Vars.lift lft1 x) (Vars.lift lft2 y) in - let eq_fix (a,b) (c,d) = f_equal (Constr.mkFix a, b) (Constr.mkFix c, d) in +let equal_stacks sigma (x, l) (y, l') = + let open EConstr in + let f_equal (x,lft1) (y,lft2) = eq_constr sigma (Vars.lift lft1 x) (Vars.lift lft2 y) in + let eq_fix (a,b) (c,d) = f_equal (mkFix a, b) (mkFix c, d) in match Stack.equal f_equal eq_fix l l' with | None -> false | Some (lft1,lft2) -> f_equal (x, lft1) (y, lft2) let rec whd_state_gen ?csts ~refold ~tactic_mode flags env sigma = + let open EConstr in let open Context.Named.Declaration in - let rec whrec cst_l (x, stack as s) = + let rec whrec cst_l (x, stack) = let () = if !debug_RAKAM then let open Pp in + let pr c = Termops.print_constr (EConstr.Unsafe.to_constr c) in Feedback.msg_notice - (h 0 (str "<<" ++ Termops.print_constr x ++ + (h 0 (str "<<" ++ pr x ++ str "|" ++ cut () ++ Cst_stack.pr cst_l ++ - str "|" ++ cut () ++ Stack.pr Termops.print_constr stack ++ + str "|" ++ cut () ++ Stack.pr pr stack ++ str ">>")) in + let c0 = EConstr.kind sigma x in let fold () = let () = if !debug_RAKAM then let open Pp in Feedback.msg_notice (str "<><><><><>") in - (s,cst_l) + ((EConstr.of_kind c0, stack),cst_l) in - match kind_of_term x with + match c0 with | Rel n when CClosure.RedFlags.red_set flags CClosure.RedFlags.fDELTA -> (match lookup_rel n env with - | LocalDef (_,body,_) -> whrec Cst_stack.empty (lift n body, stack) + | LocalDef (_,body,_) -> whrec Cst_stack.empty (EConstr.of_constr (lift n body), stack) | _ -> fold ()) | Var id when CClosure.RedFlags.red_set flags (CClosure.RedFlags.fVAR id) -> (match lookup_named id env with | LocalDef (_,body,_) -> - whrec (if refold then Cst_stack.add_cst (mkVar id) cst_l else cst_l) (body, stack) + whrec (if refold then Cst_stack.add_cst (mkVar id) cst_l else cst_l) (EConstr.of_constr body, stack) | _ -> fold ()) - | Evar ev -> - (match safe_evar_value sigma ev with - | Some body -> whrec cst_l (body, stack) - | None -> fold ()) + | Evar ev -> fold () | Meta ev -> (match safe_meta_value sigma ev with - | Some body -> whrec cst_l (body, stack) + | Some body -> whrec cst_l (EConstr.of_constr body, stack) | None -> fold ()) | Const (c,u as const) when CClosure.RedFlags.red_set flags (CClosure.RedFlags.fCONST c) -> (match constant_opt_value_in env const with | None -> fold () | Some body -> + let body = EConstr.of_constr body in if not tactic_mode then whrec (if refold then Cst_stack.add_cst (mkConstU const) cst_l else cst_l) (body, stack) @@ -863,12 +892,12 @@ let rec whd_state_gen ?csts ~refold ~tactic_mode flags env sigma = let (tm',sk'),cst_l' = whrec (Cst_stack.add_cst (mkConstU const) cst_l) (body, app_sk) in - let rec is_case x = match kind_of_term x with + let rec is_case x = match EConstr.kind sigma x with | Lambda (_,_, x) | LetIn (_,_,_, x) | Cast (x, _,_) -> is_case x | App (hd, _) -> is_case hd | Case _ -> true | _ -> false in - if equal_stacks (x, app_sk) (tm', sk') + if equal_stacks sigma (x, app_sk) (tm', sk') || Stack.will_expose_iota sk' || is_case tm' then fold () @@ -896,7 +925,7 @@ let rec whd_state_gen ?csts ~refold ~tactic_mode flags env sigma = | None -> let stack' = (c, Stack.Proj (npars, arg, p, cst_l) :: stack) in let stack'', csts = whrec Cst_stack.empty stack' in - if equal_stacks stack' stack'' then fold () + if equal_stacks sigma stack' stack'' then fold () else stack'', csts | Some (recargs, nargs, flags) -> if (List.mem `ReductionNeverUnfold flags @@ -926,7 +955,7 @@ let rec whd_state_gen ?csts ~refold ~tactic_mode flags env sigma = Stack.append_app [|c|] bef,cst_l)::s')) | LetIn (_,b,_,c) when CClosure.RedFlags.red_set flags CClosure.RedFlags.fZETA -> - apply_subst whrec [b] refold cst_l c stack + apply_subst (fun _ -> whrec) [b] sigma refold cst_l c stack | Cast (c,_,_) -> whrec cst_l (c, stack) | App (f,cl) -> whrec @@ -935,20 +964,20 @@ let rec whd_state_gen ?csts ~refold ~tactic_mode flags env sigma = | Lambda (na,t,c) -> (match Stack.decomp stack with | Some _ when CClosure.RedFlags.red_set flags CClosure.RedFlags.fBETA -> - apply_subst whrec [] refold cst_l x stack + apply_subst (fun _ -> whrec) [] sigma refold cst_l x stack | None when CClosure.RedFlags.red_set flags CClosure.RedFlags.fETA -> - let env' = push_rel (LocalAssum (na,t)) env in + let env' = push_rel (local_assum (na, t)) env in let whrec' = whd_state_gen ~refold ~tactic_mode flags env' sigma in - (match kind_of_term (Stack.zip ~refold (fst (whrec' (c, Stack.empty)))) with + (match EConstr.kind sigma (Stack.zip ~refold sigma (fst (whrec' (c, Stack.empty)))) with | App (f,cl) -> let napp = Array.length cl in if napp > 0 then let (x', l'),_ = whrec' (Array.last cl, Stack.empty) in - match kind_of_term x', l' with + match EConstr.kind sigma x', l' with | Rel 1, [] -> let lc = Array.sub cl 0 (napp-1) in - let u = if Int.equal napp 1 then f else appvect (f,lc) in - if noccurn 1 u then (pop (EConstr.of_constr u),Stack.empty),Cst_stack.empty else fold () + let u = if Int.equal napp 1 then f else mkApp (f,lc) in + if Vars.noccurn sigma 1 u then (EConstr.of_constr (pop u),Stack.empty),Cst_stack.empty else fold () | _ -> fold () else fold () | _ -> fold ()) @@ -973,11 +1002,11 @@ let rec whd_state_gen ?csts ~refold ~tactic_mode flags env sigma = |args, (Stack.Proj (n,m,p,_)::s') when use_match -> whrec Cst_stack.empty (Stack.nth args (n+m), s') |args, (Stack.Fix (f,s',cst_l)::s'') when use_fix -> - let x' = Stack.zip(x,args) in + let x' = Stack.zip sigma (x, args) in let out_sk = s' @ (Stack.append_app [|x'|] s'') in reduce_and_refold_fix whrec env sigma refold cst_l f out_sk |args, (Stack.Cst (const,curr,remains,s',cst_l) :: s'') -> - let x' = Stack.zip(x,args) in + let x' = Stack.zip sigma (x, args) in begin match remains with | [] -> (match const with @@ -985,6 +1014,7 @@ let rec whd_state_gen ?csts ~refold ~tactic_mode flags env sigma = (match constant_opt_value_in env const with | None -> fold () | Some body -> + let body = EConstr.of_constr body in whrec (if refold then Cst_stack.add_cst (mkConstU const) cst_l else cst_l) (body, s' @ (Stack.append_app [|x'|] s''))) | Stack.Cst_proj p -> @@ -1020,31 +1050,34 @@ let rec whd_state_gen ?csts ~refold ~tactic_mode flags env sigma = in fun xs -> let (s,cst_l as res) = whrec (Option.default Cst_stack.empty csts) xs in - if tactic_mode then (Stack.best_state s cst_l,Cst_stack.empty) else res + if tactic_mode then (Stack.best_state sigma s cst_l,Cst_stack.empty) else res (** reduction machine without global env and refold machinery *) let local_whd_state_gen flags sigma = - let rec whrec (x, stack as s) = - match kind_of_term x with + let open EConstr in + let rec whrec (x, stack) = + let c0 = EConstr.kind sigma x in + let s = (EConstr.of_kind c0, stack) in + match c0 with | LetIn (_,b,_,c) when CClosure.RedFlags.red_set flags CClosure.RedFlags.fZETA -> - stacklam whrec [b] c stack + stacklam whrec [b] sigma c stack | Cast (c,_,_) -> whrec (c, stack) | App (f,cl) -> whrec (f, Stack.append_app cl stack) | Lambda (_,_,c) -> (match Stack.decomp stack with | Some (a,m) when CClosure.RedFlags.red_set flags CClosure.RedFlags.fBETA -> - stacklam whrec [a] c m + stacklam whrec [a] sigma c m | None when CClosure.RedFlags.red_set flags CClosure.RedFlags.fETA -> - (match kind_of_term (Stack.zip (whrec (c, Stack.empty))) with + (match EConstr.kind sigma (Stack.zip sigma (whrec (c, Stack.empty))) with | App (f,cl) -> let napp = Array.length cl in if napp > 0 then let x', l' = whrec (Array.last cl, Stack.empty) in - match kind_of_term x', l' with + match EConstr.kind sigma x', l' with | Rel 1, [] -> let lc = Array.sub cl 0 (napp-1) in - let u = if Int.equal napp 1 then f else appvect (f,lc) in - if noccurn 1 u then (pop (EConstr.of_constr u),Stack.empty) else s + let u = if Int.equal napp 1 then f else mkApp (f,lc) in + if Vars.noccurn sigma 1 u then (EConstr.of_constr (pop u),Stack.empty) else s | _ -> s else s | _ -> s) @@ -1064,14 +1097,10 @@ let local_whd_state_gen flags sigma = |None -> s |Some (bef,arg,s') -> whrec (arg, Stack.Fix(f,bef,Cst_stack.empty)::s')) - | Evar ev -> - (match safe_evar_value sigma ev with - Some c -> whrec (c,stack) - | None -> s) - + | Evar ev -> s | Meta ev -> (match safe_meta_value sigma ev with - Some c -> whrec (c,stack) + Some c -> whrec (EConstr.of_constr c,stack) | None -> s) | Construct ((ind,c),u) -> @@ -1084,8 +1113,8 @@ let local_whd_state_gen flags sigma = |args, (Stack.Proj (n,m,p,_) :: s') when use_match -> whrec (Stack.nth args (n+m), s') |args, (Stack.Fix (f,s',cst)::s'') when use_fix -> - let x' = Stack.zip(x,args) in - whrec (contract_fix f, s' @ (Stack.append_app [|x'|] s'')) + let x' = Stack.zip sigma (x,args) in + whrec (contract_fix sigma f, s' @ (Stack.append_app [|x'|] s'')) |_, (Stack.App _|Stack.Update _|Stack.Shift _|Stack.Cst _)::_ -> assert false |_, _ -> s else s @@ -1094,7 +1123,7 @@ let local_whd_state_gen flags sigma = if CClosure.RedFlags.red_set flags CClosure.RedFlags.fCOFIX then match Stack.strip_app stack with |args, ((Stack.Case _ | Stack.Proj _)::s') -> - whrec (contract_cofix cofix, stack) + whrec (contract_cofix sigma cofix, stack) |_ -> s else s @@ -1107,7 +1136,7 @@ let raw_whd_state_gen flags env = f let stack_red_of_state_red f = - let f sigma x = decompose_app (Stack.zip (f sigma (x, Stack.empty))) in + let f sigma x = EConstr.decompose_app sigma (Stack.zip sigma (f sigma (x, Stack.empty))) in f (* Drops the Cst_stack *) @@ -1115,11 +1144,11 @@ let iterate_whd_gen refold flags env sigma s = let rec aux t = let (hd,sk),_ = whd_state_gen refold false flags env sigma (t,Stack.empty) in let whd_sk = Stack.map aux sk in - Stack.zip ~refold (hd,whd_sk) + Stack.zip sigma ~refold (hd,whd_sk) in aux s let red_of_state_red f sigma x = - Stack.zip (f sigma (x,Stack.empty)) + EConstr.Unsafe.to_constr (Stack.zip sigma (f sigma (x,Stack.empty))) (* 0. No Reduction Functions *) @@ -1174,7 +1203,7 @@ let whd_allnolet env = (* 4. Ad-hoc eta reduction, does not subsitute evars *) -let shrink_eta c = Stack.zip (local_whd_state_gen eta Evd.empty (c,Stack.empty)) +let shrink_eta c = EConstr.Unsafe.to_constr (Stack.zip Evd.empty (local_whd_state_gen eta Evd.empty (c,Stack.empty))) (* 5. Zeta Reduction Functions *) @@ -1198,7 +1227,7 @@ let clos_norm_flags flgs env sigma t = let evars ev = safe_evar_value sigma ev in CClosure.norm_val (CClosure.create_clos_infos ~evars flgs env) - (CClosure.inject t) + (CClosure.inject (EConstr.Unsafe.to_constr t)) with e when is_anomaly e -> error "Tried to normalize ill-typed term" let nf_beta = clos_norm_flags CClosure.beta (Global.env ()) @@ -1239,7 +1268,15 @@ let report_anomaly _ = let e = CErrors.push e in iraise e -let test_trans_conversion (f: constr Reduction.extended_conversion_function) reds env sigma x y = +let f_conv ?l2r ?reds env ?evars x y = + let inj = EConstr.Unsafe.to_constr in + Reduction.conv ?l2r ?reds env ?evars (inj x) (inj y) + +let f_conv_leq ?l2r ?reds env ?evars x y = + let inj = EConstr.Unsafe.to_constr in + Reduction.conv_leq ?l2r ?reds env ?evars (inj x) (inj y) + +let test_trans_conversion (f: EConstr.t Reduction.extended_conversion_function) reds env sigma x y = try let evars ev = safe_evar_value sigma ev in let _ = f ~reds env ~evars:(evars, Evd.universes sigma) x y in @@ -1247,16 +1284,16 @@ let test_trans_conversion (f: constr Reduction.extended_conversion_function) red with Reduction.NotConvertible -> false | e when is_anomaly e -> report_anomaly e -let is_conv ?(reds=full_transparent_state) env sigma = test_trans_conversion Reduction.conv reds env sigma -let is_conv_leq ?(reds=full_transparent_state) env sigma = test_trans_conversion Reduction.conv_leq reds env sigma +let is_conv ?(reds=full_transparent_state) env sigma = test_trans_conversion f_conv reds env sigma +let is_conv_leq ?(reds=full_transparent_state) env sigma = test_trans_conversion f_conv_leq reds env sigma let is_fconv ?(reds=full_transparent_state) = function | Reduction.CONV -> is_conv ~reds | Reduction.CUMUL -> is_conv_leq ~reds let check_conv ?(pb=Reduction.CUMUL) ?(ts=full_transparent_state) env sigma x y = let f = match pb with - | Reduction.CONV -> Reduction.conv - | Reduction.CUMUL -> Reduction.conv_leq + | Reduction.CONV -> f_conv + | Reduction.CUMUL -> f_conv_leq in try f ~reds:ts env ~evars:(safe_evar_value sigma, Evd.universes sigma) x y; true with Reduction.NotConvertible -> false @@ -1320,37 +1357,38 @@ let vm_infer_conv ?(pb=Reduction.CUMUL) env t1 t2 = (* Special-Purpose Reduction *) (********************************************************************) -let whd_meta sigma c = match kind_of_term c with - | Meta p -> (try meta_value sigma p with Not_found -> c) - | _ -> c +let whd_meta sigma c = match EConstr.kind sigma c with + | Meta p -> (try meta_value sigma p with Not_found -> EConstr.Unsafe.to_constr c) + | _ -> EConstr.Unsafe.to_constr c let default_plain_instance_ident = Id.of_string "H" (* Try to replace all metas. Does not replace metas in the metas' values * Differs from (strong whd_meta). *) -let plain_instance s c = - let rec irec n u = match kind_of_term u with - | Meta p -> (try lift n (Metamap.find p s) with Not_found -> u) - | App (f,l) when isCast f -> - let (f,_,t) = destCast f in +let plain_instance sigma s c = + let open EConstr in + let rec irec n u = match EConstr.kind sigma u with + | Meta p -> (try Vars.lift n (Metamap.find p s) with Not_found -> u) + | App (f,l) when isCast sigma f -> + let (f,_,t) = destCast sigma f in let l' = CArray.Fun1.smartmap irec n l in - (match kind_of_term f with + (match EConstr.kind sigma f with | Meta p -> (* Don't flatten application nodes: this is used to extract a proof-term from a proof-tree and we want to keep the structure of the proof-tree *) (try let g = Metamap.find p s in - match kind_of_term g with + match EConstr.kind sigma g with | App _ -> - let l' = CArray.Fun1.smartmap lift 1 l' in + let l' = CArray.Fun1.smartmap Vars.lift 1 l' in mkLetIn (Name default_plain_instance_ident,g,t,mkApp(mkRel 1, l')) | _ -> mkApp (g,l') with Not_found -> mkApp (f,l')) | _ -> mkApp (irec n f,l')) - | Cast (m,_,_) when isMeta m -> - (try lift n (Metamap.find (destMeta m) s) with Not_found -> u) + | Cast (m,_,_) when isMeta sigma m -> + (try Vars.lift n (Metamap.find (destMeta sigma m) s) with Not_found -> u) | _ -> - map_constr_with_binders succ irec n u + map_with_binders sigma succ irec n u in if Metamap.is_empty s then c else irec 0 c @@ -1391,7 +1429,7 @@ let plain_instance s c = let instance sigma s c = (* if s = [] then c else *) - local_strong whd_betaiota sigma (plain_instance s c) + local_strong whd_betaiota sigma (plain_instance sigma s c) (* pseudo-reduction rule: * [hnf_prod_app env s (Prod(_,B)) N --> B[N] @@ -1400,34 +1438,40 @@ let instance sigma s c = * error message. *) let hnf_prod_app env sigma t n = - match kind_of_term (whd_all env sigma t) with - | Prod (_,_,b) -> subst1 n b + let open EConstr in + match EConstr.kind sigma (EConstr.of_constr (whd_all env sigma t)) with + | Prod (_,_,b) -> EConstr.Unsafe.to_constr (Vars.subst1 n b) | _ -> anomaly ~label:"hnf_prod_app" (Pp.str "Need a product") let hnf_prod_appvect env sigma t nl = - Array.fold_left (hnf_prod_app env sigma) t nl + Array.fold_left (fun acc t -> hnf_prod_app env sigma (EConstr.of_constr acc) t) (EConstr.Unsafe.to_constr t) nl let hnf_prod_applist env sigma t nl = - List.fold_left (hnf_prod_app env sigma) t nl + List.fold_left (fun acc t -> hnf_prod_app env sigma (EConstr.of_constr acc) t) (EConstr.Unsafe.to_constr t) nl let hnf_lam_app env sigma t n = - match kind_of_term (whd_all env sigma t) with - | Lambda (_,_,b) -> subst1 n b + let open EConstr in + match EConstr.kind sigma (EConstr.of_constr (whd_all env sigma t)) with + | Lambda (_,_,b) -> EConstr.Unsafe.to_constr (Vars.subst1 n b) | _ -> anomaly ~label:"hnf_lam_app" (Pp.str "Need an abstraction") let hnf_lam_appvect env sigma t nl = - Array.fold_left (hnf_lam_app env sigma) t nl + Array.fold_left (fun acc t -> hnf_lam_app env sigma (EConstr.of_constr acc) t) (EConstr.Unsafe.to_constr t) nl let hnf_lam_applist env sigma t nl = - List.fold_left (hnf_lam_app env sigma) t nl + List.fold_left (fun acc t -> hnf_lam_app env sigma (EConstr.of_constr acc) t) (EConstr.Unsafe.to_constr t) nl + +let bind_assum (na, t) = + let inj = EConstr.Unsafe.to_constr in + (na, inj t) let splay_prod env sigma = let rec decrec env m c = let t = whd_all env sigma c in - match kind_of_term t with + match EConstr.kind sigma (EConstr.of_constr t) with | Prod (n,a,c0) -> - decrec (push_rel (LocalAssum (n,a)) env) - ((n,a)::m) c0 + decrec (push_rel (local_assum (n,a)) env) + (bind_assum (n,a)::m) c0 | _ -> m,t in decrec env [] @@ -1435,10 +1479,10 @@ let splay_prod env sigma = let splay_lam env sigma = let rec decrec env m c = let t = whd_all env sigma c in - match kind_of_term t with + match EConstr.kind sigma (EConstr.of_constr t) with | Lambda (n,a,c0) -> - decrec (push_rel (LocalAssum (n,a)) env) - ((n,a)::m) c0 + decrec (push_rel (local_assum (n,a)) env) + (bind_assum (n,a)::m) c0 | _ -> m,t in decrec env [] @@ -1446,51 +1490,51 @@ let splay_lam env sigma = let splay_prod_assum env sigma = let rec prodec_rec env l c = let t = whd_allnolet env sigma c in - match kind_of_term t with + match EConstr.kind sigma (EConstr.of_constr t) with | Prod (x,t,c) -> - prodec_rec (push_rel (LocalAssum (x,t)) env) - (Context.Rel.add (LocalAssum (x,t)) l) c + prodec_rec (push_rel (local_assum (x,t)) env) + (Context.Rel.add (local_assum (x,t)) l) c | LetIn (x,b,t,c) -> - prodec_rec (push_rel (LocalDef (x,b,t)) env) - (Context.Rel.add (LocalDef (x,b,t)) l) c + prodec_rec (push_rel (local_def (x,b,t)) env) + (Context.Rel.add (local_def (x,b,t)) l) c | Cast (c,_,_) -> prodec_rec env l c | _ -> - let t' = whd_all env sigma t in + let t' = whd_all env sigma (EConstr.of_constr t) in if Term.eq_constr t t' then l,t - else prodec_rec env l t' + else prodec_rec env l (EConstr.of_constr t') in prodec_rec env Context.Rel.empty let splay_arity env sigma c = let l, c = splay_prod env sigma c in - match kind_of_term c with + match EConstr.kind sigma (EConstr.of_constr c) with | Sort s -> l,s | _ -> invalid_arg "splay_arity" let sort_of_arity env sigma c = snd (splay_arity env sigma c) let splay_prod_n env sigma n = - let rec decrec env m ln c = if Int.equal m 0 then (ln,c) else - match kind_of_term (whd_all env sigma c) with + let rec decrec env m ln c = if Int.equal m 0 then (ln, EConstr.Unsafe.to_constr c) else + match EConstr.kind sigma (EConstr.of_constr (whd_all env sigma c)) with | Prod (n,a,c0) -> - decrec (push_rel (LocalAssum (n,a)) env) - (m-1) (Context.Rel.add (LocalAssum (n,a)) ln) c0 + decrec (push_rel (local_assum (n,a)) env) + (m-1) (Context.Rel.add (local_assum (n,a)) ln) c0 | _ -> invalid_arg "splay_prod_n" in decrec env n Context.Rel.empty let splay_lam_n env sigma n = - let rec decrec env m ln c = if Int.equal m 0 then (ln,c) else - match kind_of_term (whd_all env sigma c) with + let rec decrec env m ln c = if Int.equal m 0 then (ln, EConstr.Unsafe.to_constr c) else + match EConstr.kind sigma (EConstr.of_constr (whd_all env sigma c)) with | Lambda (n,a,c0) -> - decrec (push_rel (LocalAssum (n,a)) env) - (m-1) (Context.Rel.add (LocalAssum (n,a)) ln) c0 + decrec (push_rel (local_assum (n,a)) env) + (m-1) (Context.Rel.add (local_assum (n,a)) ln) c0 | _ -> invalid_arg "splay_lam_n" in decrec env n Context.Rel.empty let is_sort env sigma t = - match kind_of_term (whd_all env sigma t) with + match EConstr.kind sigma (EConstr.of_constr (whd_all env sigma t)) with | Sort s -> true | _ -> false @@ -1498,6 +1542,7 @@ let is_sort env sigma t = of case/fix (heuristic used by evar_conv) *) let whd_betaiota_deltazeta_for_iota_state ts env sigma csts s = + let open EConstr in let refold = get_refolding_in_reduction () in let tactic_mode = false in let rec whrec csts s = @@ -1506,15 +1551,15 @@ let whd_betaiota_deltazeta_for_iota_state ts env sigma csts s = |args, (Stack.Case _ :: _ as stack') -> let (t_o,stack_o),csts_o = whd_state_gen ~csts:csts' ~refold ~tactic_mode (CClosure.RedFlags.red_add_transparent CClosure.all ts) env sigma (t,args) in - if reducible_mind_case t_o then whrec csts_o (t_o, stack_o@stack') else s,csts' + if reducible_mind_case sigma t_o then whrec csts_o (t_o, stack_o@stack') else s,csts' |args, (Stack.Fix _ :: _ as stack') -> let (t_o,stack_o),csts_o = whd_state_gen ~csts:csts' ~refold ~tactic_mode (CClosure.RedFlags.red_add_transparent CClosure.all ts) env sigma (t,args) in - if isConstruct t_o then whrec csts_o (t_o, stack_o@stack') else s,csts' + if isConstruct sigma t_o then whrec csts_o (t_o, stack_o@stack') else s,csts' |args, (Stack.Proj (n,m,p,_) :: stack'') -> let (t_o,stack_o),csts_o = whd_state_gen ~csts:csts' ~refold ~tactic_mode (CClosure.RedFlags.red_add_transparent CClosure.all ts) env sigma (t,args) in - if isConstruct t_o then + if isConstruct sigma t_o then whrec Cst_stack.empty (Stack.nth stack_o (n+m), stack'') else s,csts' |_, ((Stack.App _| Stack.Shift _|Stack.Update _|Stack.Cst _) :: _|[]) -> s,csts' @@ -1523,9 +1568,9 @@ let whd_betaiota_deltazeta_for_iota_state ts env sigma csts s = let find_conclusion env sigma = let rec decrec env c = let t = whd_all env sigma c in - match kind_of_term t with - | Prod (x,t,c0) -> decrec (push_rel (LocalAssum (x,t)) env) c0 - | Lambda (x,t,c0) -> decrec (push_rel (LocalAssum (x,t)) env) c0 + match EConstr.kind sigma (EConstr.of_constr t) with + | Prod (x,t,c0) -> decrec (push_rel (local_assum (x,t)) env) c0 + | Lambda (x,t,c0) -> decrec (push_rel (local_assum (x,t)) env) c0 | t -> t in decrec env @@ -1539,11 +1584,12 @@ let is_arity env sigma c = (* Metas *) let meta_value evd mv = + let open EConstr in let rec valrec mv = match meta_opt_fvalue evd mv with | Some (b,_) -> let metas = Metamap.bind valrec b.freemetas in - instance evd metas b.rebus + EConstr.of_constr (instance evd metas (EConstr.of_constr b.rebus)) | None -> mkMeta mv in valrec mv @@ -1553,7 +1599,7 @@ let meta_instance sigma b = if Metaset.is_empty fm then b.rebus else let c_sigma = Metamap.bind (fun mv -> meta_value sigma mv) fm in - instance sigma c_sigma b.rebus + instance sigma c_sigma (EConstr.of_constr b.rebus) let nf_meta sigma c = meta_instance sigma (mk_freelisted c) @@ -1569,7 +1615,7 @@ let meta_reducible_instance evd b = in let metas = Metaset.fold fold fm Metamap.empty in let rec irec u = - let u = whd_betaiota Evd.empty u (** FIXME *) in + let u = whd_betaiota Evd.empty (EConstr.of_constr u) (** FIXME *) in match kind_of_term u with | Case (ci,p,c,bl) when EConstr.isMeta evd (EConstr.of_constr (strip_outer_cast evd (EConstr.of_constr c))) -> let m = destMeta (strip_outer_cast evd (EConstr.of_constr c)) in @@ -1615,32 +1661,31 @@ let meta_reducible_instance evd b = else irec b.rebus -let head_unfold_under_prod ts env _ c = +let head_unfold_under_prod ts env sigma c = + let open EConstr in let unfold (cst,u as cstu) = if Cpred.mem cst (snd ts) then match constant_opt_value_in env cstu with - | Some c -> c + | Some c -> EConstr.of_constr c | None -> mkConstU cstu else mkConstU cstu in let rec aux c = - match kind_of_term c with + match EConstr.kind sigma c with | Prod (n,t,c) -> mkProd (n,aux t, aux c) | _ -> - let (h,l) = decompose_app c in - match kind_of_term h with - | Const cst -> beta_applist (unfold cst,l) + let (h,l) = decompose_app_vect sigma c in + match EConstr.kind sigma (EConstr.of_constr h) with + | Const cst -> beta_app sigma (unfold cst, Array.map EConstr.of_constr l) | _ -> c in - aux c + EConstr.Unsafe.to_constr (aux c) let betazetaevar_applist sigma n c l = + let open EConstr in let rec stacklam n env t stack = - if Int.equal n 0 then applist (substl env t, stack) else - match kind_of_term t, stack with + if Int.equal n 0 then applist (Vars.substl env t, stack) else + match EConstr.kind sigma t, stack with | Lambda(_,_,c), arg::stacktl -> stacklam (n-1) (arg::env) c stacktl - | LetIn(_,b,_,c), _ -> stacklam (n-1) (substl env b::env) c stack - | Evar ev, _ -> - (match safe_evar_value sigma ev with - | Some body -> stacklam n env body stack - | None -> applist (substl env t, stack)) + | LetIn(_,b,_,c), _ -> stacklam (n-1) (Vars.substl env b::env) c stack + | Evar _, _ -> applist (Vars.substl env t, stack) | _ -> anomaly (Pp.str "Not enough lambda/let's") in - stacklam n [] c l + EConstr.Unsafe.to_constr (stacklam n [] c l) diff --git a/pretyping/reductionops.mli b/pretyping/reductionops.mli index 8dcf5c084..911dab0b6 100644 --- a/pretyping/reductionops.mli +++ b/pretyping/reductionops.mli @@ -38,6 +38,7 @@ val set_refolding_in_reduction : bool -> unit cst applied to params must convertible to term of the state applied to args *) module Cst_stack : sig + open EConstr type t val empty : t val add_param : constr -> t -> t @@ -45,12 +46,13 @@ module Cst_stack : sig val add_cst : constr -> t -> t val best_cst : t -> (constr * constr list) option val best_replace : Evd.evar_map -> constr -> t -> constr -> constr - val reference : t -> Constant.t option + val reference : Evd.evar_map -> t -> Constant.t option val pr : t -> Pp.std_ppcmds end module Stack : sig + open EConstr type 'a app_node val pr_app_node : ('a -> Pp.std_ppcmds) -> 'a app_node -> Pp.std_ppcmds @@ -63,7 +65,7 @@ module Stack : sig | App of 'a app_node | Case of case_info * 'a * 'a array * Cst_stack.t | Proj of int * int * projection * Cst_stack.t - | Fix of fixpoint * 'a t * Cst_stack.t + | Fix of ('a, 'a) pfixpoint * 'a t * Cst_stack.t | Cst of cst_member * int (** current foccussed arg *) * int list (** remaining args *) * 'a t * Cst_stack.t | Shift of int @@ -82,9 +84,9 @@ module Stack : sig val compare_shape : 'a t -> 'a t -> bool (** [fold2 f x sk1 sk2] folds [f] on any pair of term in [(sk1,sk2)]. @return the result and the lifts to apply on the terms *) - val fold2 : ('a -> Term.constr -> Term.constr -> 'a) -> 'a -> - Term.constr t -> Term.constr t -> 'a * int * int - val map : (Term.constr -> Term.constr) -> Term.constr t -> Term.constr t + val fold2 : ('a -> constr -> constr -> 'a) -> 'a -> + constr t -> constr t -> 'a * int * int + val map : ('a -> 'a) -> 'a t -> 'a t val append_app_list : 'a list -> 'a t -> 'a t (** if [strip_app s] = [(a,b)], then [s = a @ b] and [b] does not @@ -101,25 +103,25 @@ module Stack : sig val tail : int -> 'a t -> 'a t val nth : 'a t -> int -> 'a - val best_state : constr * constr t -> Cst_stack.t -> constr * constr t - val zip : ?refold:bool -> constr * constr t -> constr + val best_state : evar_map -> constr * constr t -> Cst_stack.t -> constr * constr t + val zip : ?refold:bool -> evar_map -> constr * constr t -> constr end (************************************************************************) -type state = constr * constr Stack.t +type state = EConstr.t * EConstr.t Stack.t -type contextual_reduction_function = env -> evar_map -> constr -> constr +type contextual_reduction_function = env -> evar_map -> EConstr.t -> constr type reduction_function = contextual_reduction_function -type local_reduction_function = evar_map -> constr -> constr +type local_reduction_function = evar_map -> EConstr.t -> constr -type e_reduction_function = { e_redfun : 'r. env -> 'r Sigma.t -> constr -> (constr, 'r) Sigma.sigma } +type e_reduction_function = { e_redfun : 'r. env -> 'r Sigma.t -> EConstr.t -> (constr, 'r) Sigma.sigma } type contextual_stack_reduction_function = - env -> evar_map -> constr -> constr * constr list + env -> evar_map -> EConstr.t -> EConstr.t * EConstr.t list type stack_reduction_function = contextual_stack_reduction_function type local_stack_reduction_function = - evar_map -> constr -> constr * constr list + evar_map -> EConstr.t -> EConstr.t * EConstr.t list type contextual_state_reduction_function = env -> evar_map -> state -> state @@ -137,13 +139,13 @@ val strong_prodspine : local_reduction_function -> local_reduction_function val stack_reduction_of_reduction : 'a reduction_function -> 'a state_reduction_function i*) -val stacklam : (state -> 'a) -> constr list -> constr -> constr Stack.t -> 'a +val stacklam : (state -> 'a) -> EConstr.t list -> evar_map -> EConstr.t -> EConstr.t Stack.t -> 'a val whd_state_gen : ?csts:Cst_stack.t -> refold:bool -> tactic_mode:bool -> CClosure.RedFlags.reds -> Environ.env -> Evd.evar_map -> state -> state * Cst_stack.t val iterate_whd_gen : bool -> CClosure.RedFlags.reds -> - Environ.env -> Evd.evar_map -> Term.constr -> Term.constr + Environ.env -> Evd.evar_map -> EConstr.constr -> EConstr.constr (** {6 Generic Optimized Reduction Function using Closures } *) @@ -196,46 +198,46 @@ val whd_zeta_stack : local_stack_reduction_function val whd_zeta_state : local_state_reduction_function val whd_zeta : local_reduction_function -val shrink_eta : constr -> constr +val shrink_eta : EConstr.t -> constr (** Various reduction functions *) val safe_evar_value : evar_map -> existential -> constr option -val beta_applist : constr * constr list -> constr - -val hnf_prod_app : env -> evar_map -> constr -> constr -> constr -val hnf_prod_appvect : env -> evar_map -> constr -> constr array -> constr -val hnf_prod_applist : env -> evar_map -> constr -> constr list -> constr -val hnf_lam_app : env -> evar_map -> constr -> constr -> constr -val hnf_lam_appvect : env -> evar_map -> constr -> constr array -> constr -val hnf_lam_applist : env -> evar_map -> constr -> constr list -> constr - -val splay_prod : env -> evar_map -> constr -> (Name.t * constr) list * constr -val splay_lam : env -> evar_map -> constr -> (Name.t * constr) list * constr -val splay_arity : env -> evar_map -> constr -> (Name.t * constr) list * sorts -val sort_of_arity : env -> evar_map -> constr -> sorts -val splay_prod_n : env -> evar_map -> int -> constr -> Context.Rel.t * constr -val splay_lam_n : env -> evar_map -> int -> constr -> Context.Rel.t * constr +val beta_applist : evar_map -> EConstr.t * EConstr.t list -> constr + +val hnf_prod_app : env -> evar_map -> EConstr.t -> EConstr.t -> constr +val hnf_prod_appvect : env -> evar_map -> EConstr.t -> EConstr.t array -> constr +val hnf_prod_applist : env -> evar_map -> EConstr.t -> EConstr.t list -> constr +val hnf_lam_app : env -> evar_map -> EConstr.t -> EConstr.t -> constr +val hnf_lam_appvect : env -> evar_map -> EConstr.t -> EConstr.t array -> constr +val hnf_lam_applist : env -> evar_map -> EConstr.t -> EConstr.t list -> constr + +val splay_prod : env -> evar_map -> EConstr.t -> (Name.t * constr) list * constr +val splay_lam : env -> evar_map -> EConstr.t -> (Name.t * constr) list * constr +val splay_arity : env -> evar_map -> EConstr.t -> (Name.t * constr) list * sorts +val sort_of_arity : env -> evar_map -> EConstr.t -> sorts +val splay_prod_n : env -> evar_map -> int -> EConstr.t -> Context.Rel.t * constr +val splay_lam_n : env -> evar_map -> int -> EConstr.t -> Context.Rel.t * constr val splay_prod_assum : - env -> evar_map -> constr -> Context.Rel.t * constr + env -> evar_map -> EConstr.t -> Context.Rel.t * constr type 'a miota_args = { - mP : constr; (** the result type *) - mconstr : constr; (** the constructor *) + mP : EConstr.t; (** the result type *) + mconstr : EConstr.t; (** the constructor *) mci : case_info; (** special info to re-build pattern *) mcargs : 'a list; (** the constructor's arguments *) mlf : 'a array } (** the branch code vector *) -val reducible_mind_case : constr -> bool -val reduce_mind_case : constr miota_args -> constr +val reducible_mind_case : evar_map -> EConstr.t -> bool +val reduce_mind_case : evar_map -> EConstr.t miota_args -> EConstr.t -val find_conclusion : env -> evar_map -> constr -> (constr,constr) kind_of_term -val is_arity : env -> evar_map -> constr -> bool -val is_sort : env -> evar_map -> types -> bool +val find_conclusion : env -> evar_map -> EConstr.t -> (EConstr.t,EConstr.t) kind_of_term +val is_arity : env -> evar_map -> EConstr.t -> bool +val is_sort : env -> evar_map -> EConstr.types -> bool -val contract_fix : ?env:Environ.env -> ?reference:Constant.t -> fixpoint -> constr -val fix_recarg : fixpoint -> constr Stack.t -> (int * constr) option +val contract_fix : ?env:Environ.env -> evar_map -> ?reference:Constant.t -> (EConstr.t, EConstr.t) pfixpoint -> EConstr.t +val fix_recarg : ('a, 'a) pfixpoint -> 'b Stack.t -> (int * 'b) option (** {6 Querying the kernel conversion oracle: opaque/transparent constants } *) val is_transparent : Environ.env -> constant tableKey -> bool @@ -247,14 +249,14 @@ type conversion_test = constraints -> constraints val pb_is_equal : conv_pb -> bool val pb_equal : conv_pb -> conv_pb -val is_conv : ?reds:transparent_state -> env -> evar_map -> constr -> constr -> bool -val is_conv_leq : ?reds:transparent_state -> env -> evar_map -> constr -> constr -> bool -val is_fconv : ?reds:transparent_state -> conv_pb -> env -> evar_map -> constr -> constr -> bool +val is_conv : ?reds:transparent_state -> env -> evar_map -> EConstr.t -> EConstr.t -> bool +val is_conv_leq : ?reds:transparent_state -> env -> evar_map -> EConstr.t -> EConstr.t -> bool +val is_fconv : ?reds:transparent_state -> conv_pb -> env -> evar_map -> EConstr.t -> EConstr.t -> bool (** [check_conv] Checks universe constraints only. pb defaults to CUMUL and ts to a full transparent state. *) -val check_conv : ?pb:conv_pb -> ?ts:transparent_state -> env -> evar_map -> constr -> constr -> bool +val check_conv : ?pb:conv_pb -> ?ts:transparent_state -> env -> evar_map -> EConstr.t -> EConstr.t -> bool (** [infer_conv] Adds necessary universe constraints to the evar map. pb defaults to CUMUL and ts to a full transparent state. @@ -280,11 +282,11 @@ val infer_conv_gen : (conv_pb -> l2r:bool -> evar_map -> transparent_state -> (** {6 Special-Purpose Reduction Functions } *) -val whd_meta : evar_map -> constr -> constr -val plain_instance : constr Metamap.t -> constr -> constr -val instance : evar_map -> constr Metamap.t -> constr -> constr +val whd_meta : local_reduction_function +val plain_instance : evar_map -> EConstr.t Metamap.t -> EConstr.t -> EConstr.t +val instance : evar_map -> EConstr.t Metamap.t -> EConstr.t -> constr val head_unfold_under_prod : transparent_state -> reduction_function -val betazetaevar_applist : evar_map -> int -> constr -> constr list -> constr +val betazetaevar_applist : evar_map -> int -> EConstr.t -> EConstr.t list -> constr (** {6 Heuristic for Conversion with Evar } *) diff --git a/pretyping/retyping.ml b/pretyping/retyping.ml index ac3b5ef63..353bdbb89 100644 --- a/pretyping/retyping.ml +++ b/pretyping/retyping.ml @@ -53,8 +53,8 @@ let get_type_from_constraints env sigma t = if isEvar (fst (decompose_app t)) then match List.map_filter (fun (pbty,env,t1,t2) -> - if is_fconv Reduction.CONV env sigma t t1 then Some t2 - else if is_fconv Reduction.CONV env sigma t t2 then Some t1 + if is_fconv Reduction.CONV env sigma (EConstr.of_constr t) (EConstr.of_constr t1) then Some t2 + else if is_fconv Reduction.CONV env sigma (EConstr.of_constr t) (EConstr.of_constr t2) then Some t1 else None) (snd (Evd.extract_all_conv_pbs sigma)) with @@ -65,7 +65,7 @@ let get_type_from_constraints env sigma t = let rec subst_type env sigma typ = function | [] -> typ | h::rest -> - match kind_of_term (whd_all env sigma typ) with + match kind_of_term (whd_all env sigma (EConstr.of_constr typ)) with | Prod (na,c1,c2) -> subst_type env sigma (subst1 h c2) rest | _ -> retype_error NonFunctionalConstruction @@ -74,7 +74,7 @@ let rec subst_type env sigma typ = function let sort_of_atomic_type env sigma ft args = let rec concl_of_arity env n ar args = - match kind_of_term (whd_all env sigma ar), args with + match kind_of_term (whd_all env sigma (EConstr.of_constr ar)), args with | Prod (na, t, b), h::l -> concl_of_arity (push_rel (LocalDef (na, lift n h, t)) env) (n + 1) b l | Sort s, [] -> s | _ -> retype_error NotASort @@ -106,17 +106,17 @@ let retype ?(polyprop=true) sigma = | Case (_,p,c,lf) -> let Inductiveops.IndType(indf,realargs) = let t = type_of env c in - try Inductiveops.find_rectype env sigma t + try Inductiveops.find_rectype env sigma (EConstr.of_constr t) with Not_found -> try let t = get_type_from_constraints env sigma t in - Inductiveops.find_rectype env sigma t + Inductiveops.find_rectype env sigma (EConstr.of_constr t) with Not_found -> retype_error BadRecursiveType in let n = inductive_nrealdecls_env env (fst (fst (dest_ind_family indf))) in - let t = betazetaevar_applist sigma n p realargs in - (match kind_of_term (whd_all env sigma (type_of env t)) with - | Prod _ -> whd_beta sigma (applist (t, [c])) + let t = betazetaevar_applist sigma n (EConstr.of_constr p) (List.map EConstr.of_constr realargs) in + (match kind_of_term (whd_all env sigma (EConstr.of_constr (type_of env t))) with + | Prod _ -> whd_beta sigma (EConstr.of_constr (applist (t, [c]))) | _ -> t) | Lambda (name,c1,c2) -> mkProd (name, c1, type_of (push_rel (LocalAssum (name,c1)) env) c2) @@ -134,7 +134,7 @@ let retype ?(polyprop=true) sigma = | Proj (p,c) -> let ty = type_of env c in (try - Inductiveops.type_of_projection_knowing_arg env sigma p c ty + Inductiveops.type_of_projection_knowing_arg env sigma p (EConstr.of_constr c) (EConstr.of_constr ty) with Invalid_argument _ -> retype_error BadRecursiveType) | Cast (c,_, t) -> t | Sort _ | Prod _ -> mkSort (sort_of env cstr) @@ -159,7 +159,7 @@ let retype ?(polyprop=true) sigma = sort_of_atomic_type env sigma t args | App(f,args) -> sort_of_atomic_type env sigma (type_of env f) args | Lambda _ | Fix _ | Construct _ -> retype_error NotAType - | _ -> decomp_sort env sigma (type_of env t) + | _ -> decomp_sort env sigma (EConstr.of_constr (type_of env t)) and sort_family_of env t = match kind_of_term t with @@ -178,7 +178,7 @@ let retype ?(polyprop=true) sigma = family_of_sort (sort_of_atomic_type env sigma (type_of env f) args) | Lambda _ | Fix _ | Construct _ -> retype_error NotAType | _ -> - family_of_sort (decomp_sort env sigma (type_of env t)) + family_of_sort (decomp_sort env sigma (EConstr.of_constr (type_of env t))) and type_of_global_reference_knowing_parameters env c args = let argtyps = @@ -207,11 +207,10 @@ let type_of_global_reference_knowing_parameters env sigma c args = let _,_,_,f = retype sigma in anomaly_on_error (f env c) args let type_of_global_reference_knowing_conclusion env sigma c conclty = - let conclty = nf_evar sigma conclty in match kind_of_term c with | Ind (ind,u) -> let spec = Inductive.lookup_mind_specif env ind in - type_of_inductive_knowing_conclusion env sigma (spec,u) conclty + type_of_inductive_knowing_conclusion env sigma (spec,u) (EConstr.of_constr conclty) | Const cst -> let t = constant_type_in env cst in (* TODO *) @@ -251,7 +250,7 @@ let sorts_of_context env evc ctxt = let expand_projection env sigma pr c args = let ty = get_type_of ~lax:true env sigma c in let (i,u), ind_args = - try Inductiveops.find_mrectype env sigma ty + try Inductiveops.find_mrectype env sigma (EConstr.of_constr ty) with Not_found -> retype_error BadRecursiveType in mkApp (mkConstU (Projection.constant pr,u), diff --git a/pretyping/tacred.ml b/pretyping/tacred.ml index ff76abe37..357a80f44 100644 --- a/pretyping/tacred.ml +++ b/pretyping/tacred.ml @@ -88,11 +88,12 @@ let evaluable_reference_eq r1 r2 = match r1, r2 with | _ -> false let mkEvalRef ref u = + let open EConstr in match ref with | EvalConst cst -> mkConstU (cst,u) | EvalVar id -> mkVar id | EvalRel n -> mkRel n - | EvalEvar ev -> mkEvar ev + | EvalEvar ev -> EConstr.of_constr (Constr.mkEvar ev) let isEvalRef env c = match kind_of_term c with | Const (sp,_) -> is_evaluable env (EvalConstRef sp) @@ -132,18 +133,18 @@ exception NotEvaluable let reference_value env sigma c u = match reference_opt_value env sigma c u with | None -> raise NotEvaluable - | Some d -> d + | Some d -> EConstr.of_constr d (************************************************************************) (* Reduction of constants hiding a fixpoint (e.g. for "simpl" tactic). *) (* One reuses the name of the function after reduction of the fixpoint *) type constant_evaluation = - | EliminationFix of int * int * (int * (int * constr) list * int) + | EliminationFix of int * int * (int * (int * EConstr.t) list * int) | EliminationMutualFix of int * evaluable_reference * ((int*evaluable_reference) option array * - (int * (int * constr) list * int)) + (int * (int * EConstr.t) list * int)) | EliminationCases of int | EliminationProj of int | NotAnElimination @@ -184,7 +185,7 @@ let check_fix_reversibility sigma labs args ((lv,i),(_,tys,bds)) = let nbfix = Array.length bds in let li = List.map - (function d -> match kind_of_term d with + (function d -> match EConstr.kind sigma d with | Rel k -> if Array.for_all (noccurn k) tys @@ -202,7 +203,7 @@ let check_fix_reversibility sigma labs args ((lv,i),(_,tys,bds)) = raise Elimconst; List.iteri (fun i t_i -> if not (Int.List.mem_assoc (i+1) li) then - let fvs = List.map ((+) (i+1)) (Int.Set.elements (free_rels sigma (EConstr.of_constr t_i))) in + let fvs = List.map ((+) (i+1)) (Int.Set.elements (free_rels sigma t_i)) in match List.intersect Int.equal fvs reversible_rels with | [] -> () | _ -> raise Elimconst) @@ -239,11 +240,11 @@ let invert_name labs l na0 env sigma ref = function | None -> None | Some c -> let labs',ccl = decompose_lam c in - let _, l' = whd_betalet_stack sigma ccl in + let _, l' = whd_betalet_stack sigma (EConstr.of_constr ccl) in let labs' = List.map snd labs' in (** ppedrot: there used to be generic equality on terms here *) if List.equal eq_constr labs' labs && - List.equal eq_constr l l' then Some (minfxargs,ref) + List.equal (fun c1 c2 -> EConstr.eq_constr sigma c1 c2) l l' then Some (minfxargs,ref) else None with Not_found (* Undefined ref *) -> None end @@ -255,11 +256,12 @@ let invert_name labs l na0 env sigma ref = function let compute_consteval_direct env sigma ref = let rec srec env n labs onlyproj c = - let c',l = whd_betadeltazeta_stack env sigma c in + let c',l = whd_betadeltazeta_stack env sigma (EConstr.of_constr c) in + let c' = EConstr.Unsafe.to_constr c' in match kind_of_term c' with | Lambda (id,t,g) when List.is_empty l && not onlyproj -> let open Context.Rel.Declaration in - srec (push_rel (LocalAssum (id,t)) env) (n+1) (t::labs) onlyproj g + srec (push_rel (LocalAssum (id,t)) env) (n+1) (EConstr.of_constr t::labs) onlyproj g | Fix fix when not onlyproj -> (try check_fix_reversibility sigma labs l fix with Elimconst -> NotAnElimination) @@ -274,8 +276,9 @@ let compute_consteval_direct env sigma ref = let compute_consteval_mutual_fix env sigma ref = let rec srec env minarg labs ref c = - let c',l = whd_betalet_stack sigma c in + let c',l = whd_betalet_stack sigma (EConstr.of_constr c) in let nargs = List.length l in + let c' = EConstr.Unsafe.to_constr c' in match kind_of_term c' with | Lambda (na,t,g) when List.is_empty l -> let open Context.Rel.Declaration in @@ -345,6 +348,7 @@ let reference_eval env sigma = function let x = Name default_dependent_ident let make_elim_fun (names,(nbfix,lv,n)) u largs = + let open EConstr in let lu = List.firstn n largs in let p = List.length lv in let lyi = List.map fst lv in @@ -353,17 +357,17 @@ let make_elim_fun (names,(nbfix,lv,n)) u largs = (* k from the comment is q+1 *) try mkRel (p+1-(List.index Int.equal (n-q) lyi)) with Not_found -> aq) - 0 (List.map (lift p) lu) + 0 (List.map (Vars.lift p) lu) in fun i -> match names.(i) with | None -> None | Some (minargs,ref) -> - let body = applistc (mkEvalRef ref u) la in + let body = applist (mkEvalRef ref u, la) in let g = List.fold_left_i (fun q (* j = n+1-q *) c (ij,tij) -> - let subst = List.map (lift (-q)) (List.firstn (n-ij) la) in - let tij' = substl (List.rev subst) tij in + let subst = List.map (Vars.lift (-q)) (List.firstn (n-ij) la) in + let tij' = Vars.substl (List.rev subst) tij in mkLambda (x,tij',c)) 1 body (List.rev lv) in Some (minargs,g) @@ -380,10 +384,11 @@ let venv = let open Context.Named.Declaration in as a problem variable: an evar that can be instantiated either by vfx (expanded fixpoint) or vfun (named function). *) let substl_with_function subst sigma constr = + let open EConstr in let evd = ref sigma in let minargs = ref Evar.Map.empty in let v = Array.of_list subst in - let rec subst_total k c = match kind_of_term c with + let rec subst_total k c = match EConstr.kind sigma c with | Rel i when k < i -> if i <= k + Array.length v then match v.(i-k-1) with @@ -393,11 +398,11 @@ let substl_with_function subst sigma constr = let sigma = Sigma.to_evar_map sigma in evd := sigma; minargs := Evar.Map.add evk min !minargs; - lift k (mkEvar (evk, [|fx;ref|])) - | (fx, None) -> lift k fx + Vars.lift k (mkEvar (evk, [|fx;ref|])) + | (fx, None) -> Vars.lift k fx else mkRel (i - Array.length v) | _ -> - map_constr_with_binders succ subst_total k c in + map_with_binders sigma succ subst_total k c in let c = subst_total 0 constr in (c, !evd, !minargs) @@ -408,27 +413,28 @@ exception Partial let solve_arity_problem env sigma fxminargs c = let evm = ref sigma in let set_fix i = evm := Evd.define i (mkVar vfx) !evm in + let open EConstr in let rec check strict c = - let c' = whd_betaiotazeta sigma c in - let (h,rcargs) = decompose_app c' in + let c' = EConstr.of_constr (whd_betaiotazeta sigma c) in + let (h,rcargs) = decompose_app_vect sigma c' in match kind_of_term h with Evar(i,_) when Evar.Map.mem i fxminargs && not (Evd.is_defined !evm i) -> let minargs = Evar.Map.find i fxminargs in - if List.length rcargs < minargs then + if Array.length rcargs < minargs then if strict then set_fix i else raise Partial; - List.iter (check strict) rcargs + Array.iter (EConstr.of_constr %> check strict) rcargs | (Var _|Const _) when isEvalRef env h -> (let ev, u = destEvalRefU h in match reference_opt_value env sigma ev u with | Some h' -> let bak = !evm in - (try List.iter (check false) rcargs + (try Array.iter (EConstr.of_constr %> check false) rcargs with Partial -> evm := bak; - check strict (applist(h',rcargs))) - | None -> List.iter (check strict) rcargs) - | _ -> iter_constr (check strict) c' in + check strict (EConstr.of_constr (Constr.mkApp(h',rcargs)))) + | None -> Array.iter (EConstr.of_constr %> check strict) rcargs) + | _ -> EConstr.iter sigma (check strict) c' in check true c; !evm @@ -446,59 +452,62 @@ let substl_checking_arity env subst sigma c = Some c' -> c' | None -> f) | _ -> map_constr nf_fix c in - nf_fix body + EConstr.of_constr (nf_fix (EConstr.Unsafe.to_constr body)) -type fix_reduction_result = NotReducible | Reduced of (constr*constr list) +type fix_reduction_result = NotReducible | Reduced of (EConstr.t * EConstr.t list) let reduce_fix whdfun sigma fix stack = match fix_recarg fix (Stack.append_app_list stack Stack.empty) with | None -> NotReducible | Some (recargnum,recarg) -> let (recarg'hd,_ as recarg') = whdfun sigma recarg in - let stack' = List.assign stack recargnum (applist recarg') in - (match kind_of_term recarg'hd with - | Construct _ -> Reduced (contract_fix fix, stack') + let stack' = List.assign stack recargnum (EConstr.applist recarg') in + (match EConstr.kind sigma recarg'hd with + | Construct _ -> Reduced (contract_fix sigma fix, stack') | _ -> NotReducible) let contract_fix_use_function env sigma f ((recindices,bodynum),(_names,_types,bodies as typedbodies)) = + let open EConstr in let nbodies = Array.length recindices in let make_Fi j = (mkFix((recindices,j),typedbodies), f j) in let lbodies = List.init nbodies make_Fi in - substl_checking_arity env (List.rev lbodies) sigma (nf_beta sigma bodies.(bodynum)) + substl_checking_arity env (List.rev lbodies) sigma (EConstr.of_constr (nf_beta sigma bodies.(bodynum))) let reduce_fix_use_function env sigma f whfun fix stack = match fix_recarg fix (Stack.append_app_list stack Stack.empty) with | None -> NotReducible | Some (recargnum,recarg) -> let (recarg'hd,_ as recarg') = - if isRel recarg then + if EConstr.isRel sigma recarg then (* The recarg cannot be a local def, no worry about the right env *) (recarg, []) else whfun recarg in - let stack' = List.assign stack recargnum (applist recarg') in - (match kind_of_term recarg'hd with + let stack' = List.assign stack recargnum (EConstr.applist recarg') in + (match EConstr.kind sigma recarg'hd with | Construct _ -> Reduced (contract_fix_use_function env sigma f fix,stack') | _ -> NotReducible) let contract_cofix_use_function env sigma f (bodynum,(_names,_,bodies as typedbodies)) = + let open EConstr in let nbodies = Array.length bodies in let make_Fi j = (mkCoFix(j,typedbodies), f j) in let subbodies = List.init nbodies make_Fi in substl_checking_arity env (List.rev subbodies) - sigma (nf_beta sigma bodies.(bodynum)) + sigma (EConstr.of_constr (nf_beta sigma bodies.(bodynum))) let reduce_mind_case_use_function func env sigma mia = - match kind_of_term mia.mconstr with + let open EConstr in + match EConstr.kind sigma mia.mconstr with | Construct ((ind_sp,i),u) -> let real_cargs = List.skipn mia.mci.ci_npar mia.mcargs in applist (mia.mlf.(i-1), real_cargs) | CoFix (bodynum,(names,_,_) as cofix) -> let build_cofix_name = - if isConst func then + if isConst sigma func then let minargs = List.length mia.mcargs in fun i -> if Int.equal i bodynum then Some (minargs,func) @@ -510,7 +519,7 @@ let reduce_mind_case_use_function func env sigma mia = the block was indeed initially built as a global definition *) let kn = map_puniverses (fun x -> con_with_label x (Label.of_id id)) - (destConst func) + (destConst sigma func) in try match constant_opt_value_in env kn with | None -> None @@ -525,13 +534,13 @@ let reduce_mind_case_use_function func env sigma mia = | _ -> assert false -let match_eval_ref env constr = - match kind_of_term constr with +let match_eval_ref env sigma constr = + match EConstr.kind sigma constr with | Const (sp, u) when is_evaluable env (EvalConstRef sp) -> Some (EvalConst sp, u) | Var id when is_evaluable env (EvalVarRef id) -> Some (EvalVar id, Univ.Instance.empty) | Rel i -> Some (EvalRel i, Univ.Instance.empty) - | Evar ev -> Some (EvalEvar ev, Univ.Instance.empty) + | Evar (evk, args) -> Some (EvalEvar (evk, Array.map EConstr.Unsafe.to_constr args), Univ.Instance.empty) | _ -> None let match_eval_ref_value env sigma constr = @@ -548,20 +557,21 @@ let match_eval_ref_value env sigma constr = let special_red_case env sigma whfun (ci, p, c, lf) = let rec redrec s = let (constr, cargs) = whfun s in - match match_eval_ref env constr with + match match_eval_ref env sigma constr with | Some (ref, u) -> (match reference_opt_value env sigma ref u with | None -> raise Redelimination | Some gvalue -> - if reducible_mind_case gvalue then + let gvalue = EConstr.of_constr gvalue in + if reducible_mind_case sigma gvalue then reduce_mind_case_use_function constr env sigma {mP=p; mconstr=gvalue; mcargs=cargs; mci=ci; mlf=lf} else - redrec (applist(gvalue, cargs))) + redrec (EConstr.applist(gvalue, cargs))) | None -> - if reducible_mind_case constr then - reduce_mind_case + if reducible_mind_case sigma constr then + reduce_mind_case sigma {mP=p; mconstr=constr; mcargs=cargs; mci=ci; mlf=lf} else @@ -574,7 +584,7 @@ let recargs = function | EvalConst c -> ReductionBehaviour.get (ConstRef c) let reduce_projection env sigma pb (recarg'hd,stack') stack = - (match kind_of_term recarg'hd with + (match EConstr.kind sigma recarg'hd with | Construct _ -> let proj_narg = pb.Declarations.proj_npars + pb.Declarations.proj_arg @@ -582,12 +592,13 @@ let reduce_projection env sigma pb (recarg'hd,stack') stack = | _ -> NotReducible) let reduce_proj env sigma whfun whfun' c = + let open EConstr in let rec redrec s = - match kind_of_term s with + match EConstr.kind sigma s with | Proj (proj, c) -> let c' = try redrec c with Redelimination -> c in let constr, cargs = whfun c' in - (match kind_of_term constr with + (match EConstr.kind sigma constr with | Construct _ -> let proj_narg = let pb = lookup_projection proj env in @@ -604,44 +615,43 @@ let reduce_proj env sigma whfun whfun' c = let whd_nothing_for_iota env sigma s = let rec whrec (x, stack as s) = - match kind_of_term x with + match EConstr.kind sigma x with | Rel n -> let open Context.Rel.Declaration in (match lookup_rel n env with - | LocalDef (_,body,_) -> whrec (lift n body, stack) + | LocalDef (_,body,_) -> whrec (EConstr.of_constr (lift n body), stack) | _ -> s) | Var id -> let open Context.Named.Declaration in (match lookup_named id env with - | LocalDef (_,body,_) -> whrec (body, stack) + | LocalDef (_,body,_) -> whrec (EConstr.of_constr body, stack) | _ -> s) - | Evar ev -> - (try whrec (Evd.existential_value sigma ev, stack) - with Evd.NotInstantiatedEvar | Not_found -> s) + | Evar ev -> s | Meta ev -> - (try whrec (Evd.meta_value sigma ev, stack) + (try whrec (EConstr.of_constr (Evd.meta_value sigma ev), stack) with Not_found -> s) | Const const when is_transparent_constant full_transparent_state (fst const) -> (match constant_opt_value_in env const with - | Some body -> whrec (body, stack) + | Some body -> whrec (EConstr.of_constr body, stack) | None -> s) - | LetIn (_,b,_,c) -> stacklam whrec [b] c stack + | LetIn (_,b,_,c) -> stacklam whrec [b] sigma c stack | Cast (c,_,_) -> whrec (c, stack) | App (f,cl) -> whrec (f, Stack.append_app cl stack) | Lambda (na,t,c) -> (match Stack.decomp stack with - | Some (a,m) -> stacklam whrec [a] c m + | Some (a,m) -> stacklam whrec [a] sigma c m | _ -> s) | x -> s in - decompose_app (Stack.zip (whrec (s,Stack.empty))) + EConstr.decompose_app sigma (Stack.zip sigma (whrec (s,Stack.empty))) (* [red_elim_const] contracts iota/fix/cofix redexes hidden behind constants by keeping the name of the constants in the recursive calls; it fails if no redex is around *) let rec red_elim_const env sigma ref u largs = + let open EConstr in let nargs = List.length largs in let largs, unfold_anyway, unfold_nonelim, nocase = match recargs ref with @@ -660,7 +670,7 @@ let rec red_elim_const env sigma ref u largs = let c = reference_value env sigma ref u in let c', lrest = whd_nothing_for_iota env sigma (applist(c,largs)) in let whfun = whd_simpl_stack env sigma in - (special_red_case env sigma whfun (destCase c'), lrest), nocase + (special_red_case env sigma whfun (EConstr.destCase sigma c'), lrest), nocase | EliminationProj n when nargs >= n -> let c = reference_value env sigma ref u in let c', lrest = whd_nothing_for_iota env sigma (applist(c,largs)) in @@ -672,9 +682,9 @@ let rec red_elim_const env sigma ref u largs = let d, lrest = whd_nothing_for_iota env sigma (applist(c,largs)) in let f = make_elim_fun ([|Some (minfxargs,ref)|],infos) u largs in let whfun = whd_construct_stack env sigma in - (match reduce_fix_use_function env sigma f whfun (destFix d) lrest with + (match reduce_fix_use_function env sigma f whfun (destFix sigma d) lrest with | NotReducible -> raise Redelimination - | Reduced (c,rest) -> (nf_beta sigma c, rest), nocase) + | Reduced (c,rest) -> (EConstr.of_constr (nf_beta sigma c), rest), nocase) | EliminationMutualFix (min,refgoal,refinfos) when nargs >= min -> let rec descend (ref,u) args = let c = reference_value env sigma ref u in @@ -682,21 +692,21 @@ let rec red_elim_const env sigma ref u largs = (c,args) else let c', lrest = whd_betalet_stack sigma (applist(c,args)) in - descend (destEvalRefU c') lrest in + descend (destEvalRefU (EConstr.Unsafe.to_constr c')) lrest in let (_, midargs as s) = descend (ref,u) largs in let d, lrest = whd_nothing_for_iota env sigma (applist s) in let f = make_elim_fun refinfos u midargs in let whfun = whd_construct_stack env sigma in - (match reduce_fix_use_function env sigma f whfun (destFix d) lrest with + (match reduce_fix_use_function env sigma f whfun (destFix sigma d) lrest with | NotReducible -> raise Redelimination - | Reduced (c,rest) -> (nf_beta sigma c, rest), nocase) + | Reduced (c,rest) -> (EConstr.of_constr (nf_beta sigma c), rest), nocase) | NotAnElimination when unfold_nonelim -> let c = reference_value env sigma ref u in - (whd_betaiotazeta sigma (applist (c, largs)), []), nocase + (EConstr.of_constr (whd_betaiotazeta sigma (applist (c, largs))), []), nocase | _ -> raise Redelimination with Redelimination when unfold_anyway -> let c = reference_value env sigma ref u in - (whd_betaiotazeta sigma (applist (c, largs)), []), nocase + (EConstr.of_constr (whd_betaiotazeta sigma (applist (c, largs))), []), nocase and reduce_params env sigma stack l = let len = List.length stack in @@ -705,8 +715,8 @@ and reduce_params env sigma stack l = else let arg = List.nth stack i in let rarg = whd_construct_stack env sigma arg in - match kind_of_term (fst rarg) with - | Construct _ -> List.assign stack i (applist rarg) + match EConstr.kind sigma (fst rarg) with + | Construct _ -> List.assign stack i (EConstr.applist rarg) | _ -> raise Redelimination) stack l @@ -715,14 +725,18 @@ and reduce_params env sigma stack l = a reducible iota/fix/cofix redex (the "simpl" tactic) *) and whd_simpl_stack env sigma = + let open EConstr in let rec redrec s = - let (x, stack as s') = decompose_app s in - match kind_of_term x with + let (x, stack) = decompose_app_vect sigma s in + let stack = Array.map_to_list EConstr.of_constr stack in + let x = EConstr.of_constr x in + let s' = (x, stack) in + match EConstr.kind sigma x with | Lambda (na,t,c) -> (match stack with | [] -> s' - | a :: rest -> redrec (beta_applist (x,stack))) - | LetIn (n,b,t,c) -> redrec (applist (substl [b] c, stack)) + | a :: rest -> redrec (EConstr.of_constr (beta_applist sigma (x, stack)))) + | LetIn (n,b,t,c) -> redrec (applist (Vars.substl [b] c, stack)) | App (f,cl) -> redrec (applist(f, (Array.to_list cl)@stack)) | Cast (c,_,_) -> redrec (applist(c, stack)) | Case (ci,p,c,lf) -> @@ -762,12 +776,12 @@ and whd_simpl_stack env sigma = with Redelimination -> s') | _ -> - match match_eval_ref env x with + match match_eval_ref env sigma x with | Some (ref, u) -> (try let sapp, nocase = red_elim_const env sigma ref u stack in let hd, _ as s'' = redrec (applist(sapp)) in - let rec is_case x = match kind_of_term x with + let rec is_case x = match EConstr.kind sigma x with | Lambda (_,_, x) | LetIn (_,_,_, x) | Cast (x, _,_) -> is_case x | App (hd, _) -> is_case hd | Case _ -> true @@ -782,13 +796,14 @@ and whd_simpl_stack env sigma = (* reduce until finding an applied constructor or fail *) and whd_construct_stack env sigma s = + let open EConstr in let (constr, cargs as s') = whd_simpl_stack env sigma s in - if reducible_mind_case constr then s' - else match match_eval_ref env constr with + if reducible_mind_case sigma constr then s' + else match match_eval_ref env sigma constr with | Some (ref, u) -> (match reference_opt_value env sigma ref u with | None -> raise Redelimination - | Some gvalue -> whd_construct_stack env sigma (applist(gvalue, cargs))) + | Some gvalue -> whd_construct_stack env sigma (applist(EConstr.of_constr gvalue, cargs))) | _ -> raise Redelimination (************************************************************************) @@ -800,12 +815,14 @@ and whd_construct_stack env sigma s = *) let try_red_product env sigma c = - let simpfun = clos_norm_flags betaiotazeta env sigma in + let simpfun c = EConstr.of_constr (clos_norm_flags betaiotazeta env sigma c) in + let inj = EConstr.Unsafe.to_constr in + let open EConstr in let rec redrec env x = - let x = whd_betaiota sigma x in - match kind_of_term x with + let x = EConstr.of_constr (whd_betaiota sigma x) in + match EConstr.kind sigma x with | App (f,l) -> - (match kind_of_term f with + (match EConstr.kind sigma f with | Fix fix -> let stack = Stack.append_app l Stack.empty in (match fix_recarg fix stack with @@ -813,17 +830,17 @@ let try_red_product env sigma c = | Some (recargnum,recarg) -> let recarg' = redrec env recarg in let stack' = Stack.assign stack recargnum recarg' in - simpfun (Stack.zip (f,stack'))) - | _ -> simpfun (appvect (redrec env f, l))) + simpfun (Stack.zip sigma (f,stack'))) + | _ -> simpfun (mkApp (redrec env f, l))) | Cast (c,_,_) -> redrec env c | Prod (x,a,b) -> let open Context.Rel.Declaration in - mkProd (x, a, redrec (push_rel (LocalAssum (x,a)) env) b) - | LetIn (x,a,b,t) -> redrec env (subst1 a t) + mkProd (x, a, redrec (push_rel (LocalAssum (x, inj a)) env) b) + | LetIn (x,a,b,t) -> redrec env (Vars.subst1 a t) | Case (ci,p,d,lf) -> simpfun (mkCase (ci,p,redrec env d,lf)) | Proj (p, c) -> let c' = - match kind_of_term c with + match EConstr.kind sigma c with | Construct _ -> c | _ -> redrec env c in @@ -832,15 +849,15 @@ let try_red_product env sigma c = | Reduced s -> simpfun (applist s) | NotReducible -> raise Redelimination) | _ -> - (match match_eval_ref env x with + (match match_eval_ref env sigma x with | Some (ref, u) -> (* TO DO: re-fold fixpoints after expansion *) (* to get true one-step reductions *) (match reference_opt_value env sigma ref u with | None -> raise Redelimination - | Some c -> c) + | Some c -> EConstr.of_constr c) | _ -> raise Redelimination) - in redrec env c + in EConstr.Unsafe.to_constr (redrec env c) let red_product env sigma c = try try_red_product env sigma c @@ -911,22 +928,23 @@ let whd_simpl_stack = immediately hides a non reducible fix or a cofix *) let whd_simpl_orelse_delta_but_fix env sigma c = + let open EConstr in let rec redrec s = let (constr, stack as s') = whd_simpl_stack env sigma s in - match match_eval_ref_value env sigma constr with + match match_eval_ref_value env sigma (EConstr.Unsafe.to_constr constr) with | Some c -> (match kind_of_term (strip_lam c) with | CoFix _ | Fix _ -> s' | Proj (p,t) when - (match kind_of_term constr with + (match EConstr.kind sigma constr with | Const (c', _) -> eq_constant (Projection.constant p) c' | _ -> false) -> let pb = Environ.lookup_projection p env in if List.length stack <= pb.Declarations.proj_npars then (** Do not show the eta-expanded form *) s' - else redrec (applist (c, stack)) - | _ -> redrec (applist(c, stack))) + else redrec (applist (EConstr.of_constr c, stack)) + | _ -> redrec (applist(EConstr.of_constr c, stack))) | None -> s' in let simpfun = clos_norm_flags betaiota env sigma in @@ -937,7 +955,7 @@ let hnf_constr = whd_simpl_orelse_delta_but_fix (* The "simpl" reduction tactic *) let whd_simpl env sigma c = - applist (whd_simpl_stack env sigma c) + EConstr.Unsafe.to_constr (EConstr.applist (whd_simpl_stack env sigma c)) let simpl env sigma c = strong whd_simpl env sigma c @@ -993,7 +1011,7 @@ let e_contextually byhead (occs,c) f = { e_redfun = begin fun env sigma t -> (* Skip inner occurrences for stable counting of occurrences *) if locs != [] then ignore (traverse_below (Some (!pos-1)) envc t); - let Sigma (t, evm, _) = (f subst).e_redfun env (Sigma.Unsafe.of_evar_map !evd) t in + let Sigma (t, evm, _) = (f subst).e_redfun env (Sigma.Unsafe.of_evar_map !evd) (EConstr.of_constr t) in (evd := Sigma.to_evar_map evm; t) end else @@ -1011,7 +1029,7 @@ let e_contextually byhead (occs,c) f = { e_redfun = begin fun env sigma t -> (fun d (env,c) -> (push_rel d env,lift_pattern 1 c)) (traverse nested) envc sigma t in - let t' = traverse None (env,c) t in + let t' = traverse None (env,c) (EConstr.Unsafe.to_constr t) in if List.exists (fun o -> o >= !pos) locs then error_invalid_occurrence locs; Sigma.Unsafe.of_pair (t', !evd) end } @@ -1028,7 +1046,7 @@ let contextually byhead occs f env sigma t = * ol is the occurrence list to find. *) let match_constr_evaluable_ref sigma c evref = - match kind_of_term c, evref with + match EConstr.kind sigma c, evref with | Const (c,u), EvalConstRef c' when eq_constant c c' -> Some u | Var id, EvalVarRef id' when id_eq id id' -> Some Univ.Instance.empty | _, _ -> None @@ -1037,7 +1055,7 @@ let substlin env sigma evalref n (nowhere_except_in,locs) c = let maxocc = List.fold_right max locs 0 in let pos = ref n in assert (List.for_all (fun x -> x >= 0) locs); - let value u = value_of_evaluable_ref env evalref u in + let value u = EConstr.of_constr (value_of_evaluable_ref env evalref u) in let rec substrec () c = if nowhere_except_in && !pos > maxocc then c else @@ -1049,9 +1067,10 @@ let substlin env sigma evalref n (nowhere_except_in,locs) c = incr pos; if ok then value u else c | None -> - map_constr_with_binders_left_to_right + let self () c = EConstr.Unsafe.to_constr (substrec () (EConstr.of_constr c)) in + EConstr.of_constr (map_constr_with_binders_left_to_right (fun _ () -> ()) - substrec () c + self () (EConstr.Unsafe.to_constr c)) in let t' = substrec () c in (!pos, t') @@ -1085,39 +1104,39 @@ let unfoldoccs env sigma (occs,name) c = nf_betaiotazeta sigma uc in match occs with - | NoOccurrences -> c + | NoOccurrences -> EConstr.Unsafe.to_constr c | AllOccurrences -> unfold env sigma name c | OnlyOccurrences l -> unfo true l | AllOccurrencesBut l -> unfo false l (* Unfold reduction tactic: *) let unfoldn loccname env sigma c = - List.fold_left (fun c occname -> unfoldoccs env sigma occname c) c loccname + EConstr.Unsafe.to_constr (List.fold_left (fun c occname -> EConstr.of_constr (unfoldoccs env sigma occname c)) c loccname) (* Re-folding constants tactics: refold com in term c *) let fold_one_com com env sigma c = let rcom = - try red_product env sigma com + try red_product env sigma (EConstr.of_constr com) with Redelimination -> error "Not reducible." in (* Reason first on the beta-iota-zeta normal form of the constant as unfold produces it, so that the "unfold f; fold f" configuration works to refold fix expressions *) - let a = subst_term sigma (EConstr.of_constr (clos_norm_flags unfold_side_red env sigma rcom)) (EConstr.of_constr c) in - if not (eq_constr a c) then + let a = subst_term sigma (EConstr.of_constr (clos_norm_flags unfold_side_red env sigma (EConstr.of_constr rcom))) c in + if not (eq_constr a (EConstr.Unsafe.to_constr c)) then subst1 com a else (* Then reason on the non beta-iota-zeta form for compatibility - even if it is probably a useless configuration *) - let a = subst_term sigma (EConstr.of_constr rcom) (EConstr.of_constr c) in + let a = subst_term sigma (EConstr.of_constr rcom) c in subst1 com a let fold_commands cl env sigma c = - List.fold_right (fun com -> fold_one_com com env sigma) (List.rev cl) c + EConstr.Unsafe.to_constr (List.fold_right (fun com c -> EConstr.of_constr (fold_one_com com env sigma c)) (List.rev cl) c) (* call by value reduction functions *) let cbv_norm_flags flags env sigma t = - cbv_norm (create_cbv_infos flags env sigma) t + cbv_norm (create_cbv_infos flags env sigma) (EConstr.Unsafe.to_constr t) let cbv_beta = cbv_norm_flags beta empty_env let cbv_betaiota = cbv_norm_flags betaiota empty_env @@ -1142,7 +1161,7 @@ let abstract_scheme env (locc,a) (c, sigma) = let pattern_occs loccs_trm = { e_redfun = begin fun env sigma c -> let sigma = Sigma.to_evar_map sigma in - let abstr_trm, sigma = List.fold_right (abstract_scheme env) loccs_trm (c,sigma) in + let abstr_trm, sigma = List.fold_right (abstract_scheme env) loccs_trm (EConstr.Unsafe.to_constr c,sigma) in try let _ = Typing.unsafe_type_of env sigma abstr_trm in Sigma.Unsafe.of_pair (applist(abstr_trm, List.map snd loccs_trm), sigma) @@ -1170,7 +1189,7 @@ let check_not_primitive_record env ind = let reduce_to_ind_gen allow_product env sigma t = let rec elimrec env t l = - let t = hnf_constr env sigma t in + let t = hnf_constr env sigma (EConstr.of_constr t) in match kind_of_term (fst (decompose_app t)) with | Ind ind-> (check_privacy env ind, it_mkProd_or_LetIn t l) | Prod (n,ty,t') -> @@ -1182,7 +1201,7 @@ let reduce_to_ind_gen allow_product env sigma t = | _ -> (* Last chance: we allow to bypass the Opaque flag (as it was partially the case between V5.10 and V8.1 *) - let t' = whd_all env sigma t in + let t' = whd_all env sigma (EConstr.of_constr t) in match kind_of_term (fst (decompose_app t')) with | Ind ind-> (check_privacy env ind, it_mkProd_or_LetIn t' l) | _ -> user_err (str"Not an inductive product.") @@ -1202,14 +1221,15 @@ let find_hnf_rectype env sigma t = exception NotStepReducible let one_step_reduce env sigma c = + let open EConstr in let rec redrec (x, stack) = - match kind_of_term x with + match EConstr.kind sigma x with | Lambda (n,t,c) -> (match stack with | [] -> raise NotStepReducible - | a :: rest -> (subst1 a c, rest)) + | a :: rest -> (Vars.subst1 a c, rest)) | App (f,cl) -> redrec (f, (Array.to_list cl)@stack) - | LetIn (_,f,_,cl) -> (subst1 f cl,stack) + | LetIn (_,f,_,cl) -> (Vars.subst1 f cl,stack) | Cast (c,_,_) -> redrec (c,stack) | Case (ci,p,c,lf) -> (try @@ -1221,13 +1241,13 @@ let one_step_reduce env sigma c = | Reduced s' -> s' | NotReducible -> raise NotStepReducible with Redelimination -> raise NotStepReducible) - | _ when isEvalRef env x -> - let ref,u = destEvalRefU x in + | _ when isEvalRef env (EConstr.Unsafe.to_constr x) -> + let ref,u = destEvalRefU (EConstr.Unsafe.to_constr x) in (try fst (red_elim_const env sigma ref u stack) with Redelimination -> match reference_opt_value env sigma ref u with - | Some d -> (d, stack) + | Some d -> (EConstr.of_constr d, stack) | None -> raise NotStepReducible) | _ -> raise NotStepReducible @@ -1249,7 +1269,7 @@ let reduce_to_ref_gen allow_product env sigma ref t = else (* lazily reduces to match the head of [t] with the expected [ref] *) let rec elimrec env t l = - let c, _ = decompose_appvect (Reductionops.whd_nored sigma t) in + let c, _ = decompose_app_vect sigma (EConstr.of_constr t) in match kind_of_term c with | Prod (n,ty,t') -> if allow_product then @@ -1264,7 +1284,7 @@ let reduce_to_ref_gen allow_product env sigma ref t = else raise Not_found with Not_found -> try - let t' = nf_betaiota sigma (one_step_reduce env sigma t) in + let t' = nf_betaiota sigma (one_step_reduce env sigma (EConstr.of_constr t)) in elimrec env t' l with NotStepReducible -> error_cannot_recognize ref in diff --git a/pretyping/typeclasses.ml b/pretyping/typeclasses.ml index 4207eccb9..50dd66ea0 100644 --- a/pretyping/typeclasses.ml +++ b/pretyping/typeclasses.ml @@ -289,7 +289,7 @@ let build_subclasses ~check env sigma glob pri = | None -> [] | Some (rels, ((tc,u), args)) -> let instapp = - Reductionops.whd_beta sigma (appvectc c (Context.Rel.to_extended_vect 0 rels)) + Reductionops.whd_beta sigma (EConstr.of_constr (appvectc c (Context.Rel.to_extended_vect 0 rels))) in let projargs = Array.of_list (args @ [instapp]) in let projs = List.map_filter diff --git a/pretyping/typing.ml b/pretyping/typing.ml index e79e3d46f..e3d1c1741 100644 --- a/pretyping/typing.ml +++ b/pretyping/typing.ml @@ -36,7 +36,7 @@ let inductive_type_knowing_parameters env (ind,u) jl = Inductive.type_of_inductive_knowing_parameters env (mspec,u) paramstyp let e_type_judgment env evdref j = - match kind_of_term (whd_all env !evdref j.uj_type) with + match kind_of_term (whd_all env !evdref (EConstr.of_constr j.uj_type)) with | Sort s -> {utj_val = j.uj_val; utj_type = s } | Evar ev -> let (evd,s) = Evardefine.define_evar_as_sort env !evdref ev in @@ -54,7 +54,7 @@ let e_judge_of_apply env evdref funj argjv = { uj_val = mkApp (j_val funj, Array.map j_val argjv); uj_type = typ } | hj::restjl -> - match kind_of_term (whd_all env !evdref typ) with + match kind_of_term (whd_all env !evdref (EConstr.of_constr typ)) with | Prod (_,c1,c2) -> if Evarconv.e_cumul env evdref hj.uj_type c1 then apply_rec (n+1) (subst1 hj.uj_val c2) restjl @@ -87,7 +87,7 @@ let e_is_correct_arity env evdref c pj ind specif params = let allowed_sorts = elim_sorts specif in let error () = error_elim_arity env ind allowed_sorts c pj None in let rec srec env pt ar = - let pt' = whd_all env !evdref pt in + let pt' = whd_all env !evdref (EConstr.of_constr pt) in match kind_of_term pt', ar with | Prod (na1,a1,t), (LocalAssum (_,a1'))::ar' -> if not (Evarconv.e_cumul env evdref a1 a1') then error (); @@ -113,12 +113,12 @@ let e_type_case_branches env evdref (ind,largs) pj c = let () = e_is_correct_arity env evdref c pj ind specif params in let lc = build_branches_type ind specif params p in let n = (snd specif).Declarations.mind_nrealdecls in - let ty = whd_betaiota !evdref (lambda_applist_assum (n+1) p (realargs@[c])) in + let ty = whd_betaiota !evdref (EConstr.of_constr (lambda_applist_assum (n+1) p (realargs@[c]))) in (lc, ty) let e_judge_of_case env evdref ci pj cj lfj = let indspec = - try find_mrectype env !evdref cj.uj_type + try find_mrectype env !evdref (EConstr.of_constr cj.uj_type) with Not_found -> error_case_not_inductive env cj in let _ = check_case_info env (fst indspec) ci in let (bty,rslty) = e_type_case_branches env evdref indspec pj cj.uj_val in @@ -139,7 +139,7 @@ let check_type_fixpoint loc env evdref lna lar vdefj = (* FIXME: might depend on the level of actual parameters!*) let check_allowed_sort env sigma ind c p = let pj = Retyping.get_judgment_of env sigma p in - let ksort = family_of_sort (sort_of_arity env sigma pj.uj_type) in + let ksort = family_of_sort (sort_of_arity env sigma (EConstr.of_constr pj.uj_type)) in let specif = Global.lookup_inductive (fst ind) in let sorts = elim_sorts specif in if not (List.exists ((==) ksort) sorts) then diff --git a/pretyping/unification.ml b/pretyping/unification.ml index 3134dac6a..ede2606da 100644 --- a/pretyping/unification.ml +++ b/pretyping/unification.ml @@ -68,7 +68,7 @@ let occur_meta_or_undefined_evar evd c = let occur_meta_evd sigma mv c = let rec occrec c = (* Note: evars are not instantiated by terms with metas *) - let c = whd_evar sigma (whd_meta sigma c) in + let c = whd_evar sigma (whd_meta sigma (EConstr.of_constr c)) in match kind_of_term c with | Meta mv' when Int.equal mv mv' -> raise Occur | _ -> Constr.iter occrec c @@ -98,7 +98,7 @@ let abstract_scheme env evd c l lname_typ = (* Precondition: resulting abstraction is expected to be of type [typ] *) let abstract_list_all env evd typ c l = - let ctxt,_ = splay_prod_n env evd (List.length l) typ in + let ctxt,_ = splay_prod_n env evd (List.length l) (EConstr.of_constr typ) in let l_with_all_occs = List.map (function a -> (LikeFirst,a)) l in let p,evd = abstract_scheme env evd c l_with_all_occs ctxt in let evd,typp = @@ -480,8 +480,8 @@ let unfold_projection env p stk = let expand_key ts env sigma = function | Some (IsKey k) -> expand_table_key env k | Some (IsProj (p, c)) -> - let red = Stack.zip (fst (whd_betaiota_deltazeta_for_iota_state ts env sigma - Cst_stack.empty (c, unfold_projection env p []))) + let red = EConstr.Unsafe.to_constr (Stack.zip sigma (fst (whd_betaiota_deltazeta_for_iota_state ts env sigma + Cst_stack.empty (EConstr.of_constr c, unfold_projection env p [])))) in if Term.eq_constr (mkProj (p, c)) red then None else Some red | None -> None @@ -576,8 +576,8 @@ let constr_cmp pb sigma flags t u = sigma, false let do_reduce ts (env, nb) sigma c = - Stack.zip (fst (whd_betaiota_deltazeta_for_iota_state - ts env sigma Cst_stack.empty (c, Stack.empty))) + EConstr.Unsafe.to_constr (Stack.zip sigma (fst (whd_betaiota_deltazeta_for_iota_state + ts env sigma Cst_stack.empty (EConstr.of_constr c, Stack.empty)))) let use_full_betaiota flags = flags.modulo_betaiota && Flags.version_strictly_greater Flags.V8_3 @@ -977,33 +977,33 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb (match expand_key flags.modulo_delta curenv sigma cf1 with | Some c -> unirec_rec curenvnb pb opt substn - (whd_betaiotazeta sigma (mkApp(c,l1))) cN + (whd_betaiotazeta sigma (EConstr.of_constr (mkApp(c,l1)))) cN | None -> (match expand_key flags.modulo_delta curenv sigma cf2 with | Some c -> unirec_rec curenvnb pb opt substn cM - (whd_betaiotazeta sigma (mkApp(c,l2))) + (whd_betaiotazeta sigma (EConstr.of_constr (mkApp(c,l2)))) | None -> error_cannot_unify curenv sigma (cM,cN))) | Some false -> (match expand_key flags.modulo_delta curenv sigma cf2 with | Some c -> unirec_rec curenvnb pb opt substn cM - (whd_betaiotazeta sigma (mkApp(c,l2))) + (whd_betaiotazeta sigma (EConstr.of_constr (mkApp(c,l2)))) | None -> (match expand_key flags.modulo_delta curenv sigma cf1 with | Some c -> unirec_rec curenvnb pb opt substn - (whd_betaiotazeta sigma (mkApp(c,l1))) cN + (whd_betaiotazeta sigma (EConstr.of_constr (mkApp(c,l1)))) cN | None -> error_cannot_unify curenv sigma (cM,cN))) and canonical_projections (curenv, _ as curenvnb) pb opt cM cN (sigma,_,_ as substn) = let f1 () = if isApp cM then - let f1l1 = whd_nored_state sigma (cM,Stack.empty) in + let f1l1 = whd_nored_state sigma (EConstr.of_constr cM,Stack.empty) in if is_open_canonical_projection curenv sigma f1l1 then - let f2l2 = whd_nored_state sigma (cN,Stack.empty) in + let f2l2 = whd_nored_state sigma (EConstr.of_constr cN,Stack.empty) in solve_canonical_projection curenvnb pb opt cM f1l1 cN f2l2 substn else error_cannot_unify (fst curenvnb) sigma (cM,cN) else error_cannot_unify (fst curenvnb) sigma (cM,cN) @@ -1017,9 +1017,9 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb else try f1 () with e when precatchable_exception e -> if isApp cN then - let f2l2 = whd_nored_state sigma (cN, Stack.empty) in + let f2l2 = whd_nored_state sigma (EConstr.of_constr cN, Stack.empty) in if is_open_canonical_projection curenv sigma f2l2 then - let f1l1 = whd_nored_state sigma (cM, Stack.empty) in + let f1l1 = whd_nored_state sigma (EConstr.of_constr cM, Stack.empty) in solve_canonical_projection curenvnb pb opt cN f2l2 cM f1l1 substn else error_cannot_unify (fst curenvnb) sigma (cM,cN) else error_cannot_unify (fst curenvnb) sigma (cM,cN) @@ -1044,13 +1044,14 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb in try let opt' = {opt with with_types = false} in + let inj = EConstr.Unsafe.to_constr in let (substn,_,_) = Reductionops.Stack.fold2 - (fun s u1 u -> unirec_rec curenvnb pb opt' s u1 (substl ks u)) + (fun s u1 u -> unirec_rec curenvnb pb opt' s (inj u1) (substl ks (inj u))) (evd,ms,es) us2 us in let (substn,_,_) = Reductionops.Stack.fold2 - (fun s u1 u -> unirec_rec curenvnb pb opt' s u1 (substl ks u)) + (fun s u1 u -> unirec_rec curenvnb pb opt' s (inj u1) (substl ks (inj u))) substn params1 params in - let (substn,_,_) = Reductionops.Stack.fold2 (unirec_rec curenvnb pb opt') substn ts ts1 in + let (substn,_,_) = Reductionops.Stack.fold2 (fun s u1 u2 -> unirec_rec curenvnb pb opt' s (inj u1) (inj u2)) substn ts ts1 in let app = mkApp (c, Array.rev_of_list ks) in (* let substn = unirec_rec curenvnb pb b false substn t cN in *) unirec_rec curenvnb pb opt' substn c1 app @@ -1206,7 +1207,7 @@ let applyHead env (type r) (evd : r Sigma.t) n c = if Int.equal n 0 then Sigma (c, evd, p) else - match kind_of_term (whd_all env (Sigma.to_evar_map evd) cty) with + match kind_of_term (whd_all env (Sigma.to_evar_map evd) (EConstr.of_constr cty)) with | Prod (_,c1,c2) -> let Sigma (evar, evd', q) = Evarutil.new_evar env evd ~src:(Loc.ghost,Evar_kinds.GoalEvar) c1 in apprec (n-1) (mkApp(c,[|evar|])) (subst1 evar c2) (p +> q) evd' @@ -1235,7 +1236,7 @@ let w_coerce_to_type env evd c cty mvty = (* inh_conv_coerce_rigid_to should have reasoned modulo reduction but there are cases where it though it was not rigid (like in fst (nat,nat)) and stops while it could have seen that it is rigid *) - let cty = Tacred.hnf_constr env evd cty in + let cty = Tacred.hnf_constr env evd (EConstr.of_constr cty) in try_to_coerce env evd c cty tycon let w_coerce env evd mv c = @@ -1246,7 +1247,7 @@ let w_coerce env evd mv c = let unify_to_type env sigma flags c status u = let sigma, c = refresh_universes (Some false) env sigma c in let t = get_type_of env sigma (nf_meta sigma c) in - let t = nf_betaiota sigma (nf_meta sigma t) in + let t = nf_betaiota sigma (EConstr.of_constr (nf_meta sigma t)) in unify_0 env sigma CUMUL flags t u let unify_type env sigma flags mv status c = @@ -1270,7 +1271,7 @@ let order_metas metas = (* Solve an equation ?n[x1=u1..xn=un] = t where ?n is an evar *) let solve_simple_evar_eqn ts env evd ev rhs = - match solve_simple_eqn (Evarconv.evar_conv_x ts) env evd (None,ev,rhs) with + match solve_simple_eqn (Evarconv.evar_conv_x ts) env evd (None,ev,EConstr.of_constr rhs) with | UnifFailure (evd,reason) -> error_cannot_unify env evd ~reason (mkEvar ev,rhs); | Success evd -> @@ -1349,7 +1350,7 @@ let w_merge env with_types flags (evd,metas,evars) = else let evd' = if occur_meta_evd evd mv c then - if isMetaOf mv (whd_all env evd c) then evd + if isMetaOf mv (whd_all env evd (EConstr.of_constr c)) then evd else error_cannot_unify env evd (mkMeta mv,c) else meta_assign mv (c,(status,TypeProcessed)) evd in @@ -1415,7 +1416,7 @@ let w_unify_meta_types env ?(flags=default_unify_flags ()) evd = types of metavars are unifiable with the types of their instances *) let head_app sigma m = - fst (whd_nored_state sigma (m, Stack.empty)) + EConstr.Unsafe.to_constr (fst (whd_nored_state sigma (EConstr.of_constr m, Stack.empty))) let check_types env flags (sigma,_,_ as subst) m n = if isEvar_or_Meta (head_app sigma m) then @@ -1577,7 +1578,7 @@ let make_pattern_test from_prefix_of_ind is_correct_type env sigma (pending,c) = (fun test -> match test.testing_state with | None -> None | Some (sigma,_,l) -> - let c = applist (nf_evar sigma (local_strong whd_meta sigma c),l) in + let c = applist (nf_evar sigma (local_strong whd_meta sigma (EConstr.of_constr c)),l) in let univs, subst = nf_univ_variables sigma in Some (sigma,subst_univs_constr subst c)) @@ -1832,7 +1833,7 @@ let w_unify_to_subterm_all env evd ?(flags=default_unify_flags ()) (op,cl) = let w_unify_to_subterm_list env evd flags hdmeta oplist t = List.fold_right (fun op (evd,l) -> - let op = whd_meta evd op in + let op = whd_meta evd (EConstr.of_constr op) in if isMeta op then if flags.allow_K_in_toplevel_higher_order_unification then (evd,op::l) else error_abstraction_over_meta env evd hdmeta (destMeta op) @@ -1905,15 +1906,16 @@ let secondOrderAbstractionAlgo dep = if dep then secondOrderDependentAbstraction else secondOrderAbstraction let w_unify2 env evd flags dep cv_pb ty1 ty2 = - let c1, oplist1 = whd_nored_stack evd ty1 in - let c2, oplist2 = whd_nored_stack evd ty2 in - match kind_of_term c1, kind_of_term c2 with + let inj = EConstr.Unsafe.to_constr in + let c1, oplist1 = whd_nored_stack evd (EConstr.of_constr ty1) in + let c2, oplist2 = whd_nored_stack evd (EConstr.of_constr ty2) in + match EConstr.kind evd c1, EConstr.kind evd c2 with | Meta p1, _ -> (* Find the predicate *) - secondOrderAbstractionAlgo dep env evd flags ty2 (p1,oplist1) + secondOrderAbstractionAlgo dep env evd flags ty2 (p1, List.map inj oplist1) | _, Meta p2 -> (* Find the predicate *) - secondOrderAbstractionAlgo dep env evd flags ty1 (p2, oplist2) + secondOrderAbstractionAlgo dep env evd flags ty1 (p2, List.map inj oplist2) | _ -> error "w_unify2" (* The unique unification algorithm works like this: If the pattern is @@ -1937,8 +1939,8 @@ let w_unify2 env evd flags dep cv_pb ty1 ty2 = convertible and first-order otherwise. But if failed if e.g. the type of Meta(1) had meta-variables in it. *) let w_unify env evd cv_pb ?(flags=default_unify_flags ()) ty1 ty2 = - let hd1,l1 = decompose_appvect (whd_nored evd ty1) in - let hd2,l2 = decompose_appvect (whd_nored evd ty2) in + let hd1,l1 = decompose_appvect (whd_nored evd (EConstr.of_constr ty1)) in + let hd2,l2 = decompose_appvect (whd_nored evd (EConstr.of_constr ty2)) in let is_empty1 = Array.is_empty l1 in let is_empty2 = Array.is_empty l2 in match kind_of_term hd1, not is_empty1, kind_of_term hd2, not is_empty2 with diff --git a/pretyping/vnorm.ml b/pretyping/vnorm.ml index 75159bf8b..8c3de7cfd 100644 --- a/pretyping/vnorm.ml +++ b/pretyping/vnorm.ml @@ -99,13 +99,13 @@ let construct_of_constr_block = construct_of_constr false let type_of_ind env (ind, u) = type_of_inductive env (Inductive.lookup_mind_specif env ind, u) -let build_branches_type env (mind,_ as _ind) mib mip u params dep p = +let build_branches_type env sigma (mind,_ as _ind) mib mip u params dep p = let rtbl = mip.mind_reloc_tbl in (* [build_one_branch i cty] construit le type de la ieme branche (commence a 0) et les lambda correspondant aux realargs *) let build_one_branch i cty = let typi = type_constructor mind mib u cty params in - let decl,indapp = Reductionops.splay_prod env Evd.empty typi in + let decl,indapp = Reductionops.splay_prod env sigma (EConstr.of_constr typi) in let decl_with_letin,_ = decompose_prod_assum typi in let ((ind,u),cargs) = find_rectype_a env indapp in let nparams = Array.length params in @@ -131,28 +131,28 @@ let build_case_type dep p realargs c = (* La fonction de normalisation *) -let rec nf_val env v t = nf_whd env (whd_val v) t +let rec nf_val env sigma v t = nf_whd env sigma (whd_val v) t -and nf_vtype env v = nf_val env v crazy_type +and nf_vtype env sigma v = nf_val env sigma v crazy_type -and nf_whd env whd typ = +and nf_whd env sigma whd typ = match whd with | Vsort s -> mkSort s | Vprod p -> - let dom = nf_vtype env (dom p) in + let dom = nf_vtype env sigma (dom p) in let name = Name (Id.of_string "x") in let vc = body_of_vfun (nb_rel env) (codom p) in - let codom = nf_vtype (push_rel (LocalAssum (name,dom)) env) vc in + let codom = nf_vtype (push_rel (LocalAssum (name,dom)) env) sigma vc in mkProd(name,dom,codom) - | Vfun f -> nf_fun env f typ - | Vfix(f,None) -> nf_fix env f - | Vfix(f,Some vargs) -> fst (nf_fix_app env f vargs) - | Vcofix(cf,_,None) -> nf_cofix env cf + | Vfun f -> nf_fun env sigma f typ + | Vfix(f,None) -> nf_fix env sigma f + | Vfix(f,Some vargs) -> fst (nf_fix_app env sigma f vargs) + | Vcofix(cf,_,None) -> nf_cofix env sigma cf | Vcofix(cf,_,Some vargs) -> - let cfd = nf_cofix env cf in + let cfd = nf_cofix env sigma cf in let i,(_,ta,_) = destCoFix cfd in let t = ta.(i) in - let _, args = nf_args env vargs t in + let _, args = nf_args env sigma vargs t in mkApp(cfd,args) | Vconstr_const n -> construct_of_constr_const env n typ @@ -165,10 +165,10 @@ and nf_whd env whd typ = | _ -> assert false else (tag, 0) in let capp,ctyp = construct_of_constr_block env tag typ in - let args = nf_bargs env b ofs ctyp in + let args = nf_bargs env sigma b ofs ctyp in mkApp(capp,args) | Vatom_stk(Aid idkey, stk) -> - constr_type_of_idkey env idkey stk + constr_type_of_idkey env sigma idkey stk | Vatom_stk(Aind ((mi,i) as ind), stk) -> let mib = Environ.lookup_mind mi env in let nb_univs = @@ -178,12 +178,12 @@ and nf_whd env whd typ = let mk u = let pind = (ind, u) in (mkIndU pind, type_of_ind env pind) in - nf_univ_args ~nb_univs mk env stk + nf_univ_args ~nb_univs mk env sigma stk | Vatom_stk(Atype u, stk) -> assert false | Vuniv_level lvl -> assert false -and nf_univ_args ~nb_univs mk env stk = +and nf_univ_args ~nb_univs mk env sigma stk = let u = if Int.equal nb_univs 0 then Univ.Instance.empty else match stk with @@ -195,9 +195,9 @@ and nf_univ_args ~nb_univs mk env stk = | _ -> assert false in let (t,ty) = mk u in - nf_stk ~from:nb_univs env t ty stk + nf_stk ~from:nb_univs env sigma t ty stk -and constr_type_of_idkey env (idkey : Vars.id_key) stk = +and constr_type_of_idkey env sigma (idkey : Vars.id_key) stk = match idkey with | ConstKey cst -> let cbody = Environ.lookup_constant cst env in @@ -208,30 +208,30 @@ and constr_type_of_idkey env (idkey : Vars.id_key) stk = let mk u = let pcst = (cst, u) in (mkConstU pcst, Typeops.type_of_constant_in env pcst) in - nf_univ_args ~nb_univs mk env stk + nf_univ_args ~nb_univs mk env sigma stk | VarKey id -> let ty = NamedDecl.get_type (lookup_named id env) in - nf_stk env (mkVar id) ty stk + nf_stk env sigma (mkVar id) ty stk | RelKey i -> let n = (nb_rel env - i) in let ty = RelDecl.get_type (lookup_rel n env) in - nf_stk env (mkRel n) (lift n ty) stk + nf_stk env sigma (mkRel n) (lift n ty) stk -and nf_stk ?from:(from=0) env c t stk = +and nf_stk ?from:(from=0) env sigma c t stk = match stk with | [] -> c | Zapp vargs :: stk -> if nargs vargs >= from then - let t, args = nf_args ~from:from env vargs t in - nf_stk env (mkApp(c,args)) t stk + let t, args = nf_args ~from:from env sigma vargs t in + nf_stk env sigma (mkApp(c,args)) t stk else let rest = from - nargs vargs in - nf_stk ~from:rest env c t stk + nf_stk ~from:rest env sigma c t stk | Zfix (f,vargs) :: stk -> assert (from = 0) ; - let fa, typ = nf_fix_app env f vargs in + let fa, typ = nf_fix_app env sigma f vargs in let _,_,codom = decompose_prod env typ in - nf_stk env (mkApp(fa,[|c|])) (subst1 c codom) stk + nf_stk env sigma (mkApp(fa,[|c|])) (subst1 c codom) stk | Zswitch sw :: stk -> assert (from = 0) ; let ((mind,_ as ind), u), allargs = find_rectype_a env t in @@ -241,34 +241,34 @@ and nf_stk ?from:(from=0) env c t stk = let pT = hnf_prod_applist env (type_of_ind env (ind,u)) (Array.to_list params) in let pT = whd_all env pT in - let dep, p = nf_predicate env (ind,u) mip params (type_of_switch sw) pT in + let dep, p = nf_predicate env sigma (ind,u) mip params (type_of_switch sw) pT in (* Calcul du type des branches *) - let btypes = build_branches_type env ind mib mip u params dep p in + let btypes = build_branches_type env sigma ind mib mip u params dep p in (* calcul des branches *) let bsw = branch_of_switch (nb_rel env) sw in let mkbranch i (n,v) = let decl,decl_with_letin,codom = btypes.(i) in - let b = nf_val (Termops.push_rels_assum decl env) v codom in + let b = nf_val (Termops.push_rels_assum decl env) sigma v codom in Termops.it_mkLambda_or_LetIn_from_no_LetIn b decl_with_letin in let branchs = Array.mapi mkbranch bsw in let tcase = build_case_type dep p realargs c in let ci = case_info sw in - nf_stk env (mkCase(ci, p, c, branchs)) tcase stk + nf_stk env sigma (mkCase(ci, p, c, branchs)) tcase stk | Zproj p :: stk -> assert (from = 0) ; let p' = Projection.make p true in - let ty = Inductiveops.type_of_projection_knowing_arg env Evd.empty p' c t in - nf_stk env (mkProj(p',c)) ty stk + let ty = Inductiveops.type_of_projection_knowing_arg env sigma p' (EConstr.of_constr c) (EConstr.of_constr t) in + nf_stk env sigma (mkProj(p',c)) ty stk -and nf_predicate env ind mip params v pT = +and nf_predicate env sigma ind mip params v pT = match whd_val v, kind_of_term pT with | Vfun f, Prod _ -> let k = nb_rel env in let vb = body_of_vfun k f in let name,dom,codom = decompose_prod env pT in let dep,body = - nf_predicate (push_rel (LocalAssum (name,dom)) env) ind mip params vb codom in + nf_predicate (push_rel (LocalAssum (name,dom)) env) sigma ind mip params vb codom in dep, mkLambda(name,dom,body) | Vfun f, _ -> let k = nb_rel env in @@ -278,33 +278,33 @@ and nf_predicate env ind mip params v pT = let rargs = Array.init n (fun i -> mkRel (n-i)) in let params = if Int.equal n 0 then params else Array.map (lift n) params in let dom = mkApp(mkIndU ind,Array.append params rargs) in - let body = nf_vtype (push_rel (LocalAssum (name,dom)) env) vb in + let body = nf_vtype (push_rel (LocalAssum (name,dom)) env) sigma vb in true, mkLambda(name,dom,body) - | _, _ -> false, nf_val env v crazy_type + | _, _ -> false, nf_val env sigma v crazy_type -and nf_args env vargs ?from:(f=0) t = +and nf_args env sigma vargs ?from:(f=0) t = let t = ref t in let len = nargs vargs - f in let args = Array.init len (fun i -> let _,dom,codom = decompose_prod env !t in - let c = nf_val env (arg vargs (f+i)) dom in + let c = nf_val env sigma (arg vargs (f+i)) dom in t := subst1 c codom; c) in !t,args -and nf_bargs env b ofs t = +and nf_bargs env sigma b ofs t = let t = ref t in let len = bsize b - ofs in let args = Array.init len (fun i -> let _,dom,codom = decompose_prod env !t in - let c = nf_val env (bfield b (i+ofs)) dom in + let c = nf_val env sigma (bfield b (i+ofs)) dom in t := subst1 c codom; c) in args -and nf_fun env f typ = +and nf_fun env sigma f typ = let k = nb_rel env in let vb = body_of_vfun k f in let name,dom,codom = @@ -314,46 +314,46 @@ and nf_fun env f typ = CErrors.anomaly (Pp.strbrk "Returned a functional value in a type not recognized as a product type.") in - let body = nf_val (push_rel (LocalAssum (name,dom)) env) vb codom in + let body = nf_val (push_rel (LocalAssum (name,dom)) env) sigma vb codom in mkLambda(name,dom,body) -and nf_fix env f = +and nf_fix env sigma f = let init = current_fix f in let rec_args = rec_args f in let k = nb_rel env in let vb, vt = reduce_fix k f in let ndef = Array.length vt in - let ft = Array.map (fun v -> nf_val env v crazy_type) vt in + let ft = Array.map (fun v -> nf_val env sigma v crazy_type) vt in let name = Array.init ndef (fun _ -> (Name (Id.of_string "Ffix"))) in (* Third argument of the tuple is ignored by push_rec_types *) let env = push_rec_types (name,ft,ft) env in (* We lift here because the types of arguments (in tt) will be evaluated in an environment where the fixpoints have been pushed *) - let norm_vb v t = nf_fun env v (lift ndef t) in + let norm_vb v t = nf_fun env sigma v (lift ndef t) in let fb = Util.Array.map2 norm_vb vb ft in mkFix ((rec_args,init),(name,ft,fb)) -and nf_fix_app env f vargs = - let fd = nf_fix env f in +and nf_fix_app env sigma f vargs = + let fd = nf_fix env sigma f in let (_,i),(_,ta,_) = destFix fd in let t = ta.(i) in - let t, args = nf_args env vargs t in + let t, args = nf_args env sigma vargs t in mkApp(fd,args),t -and nf_cofix env cf = +and nf_cofix env sigma cf = let init = current_cofix cf in let k = nb_rel env in let vb,vt = reduce_cofix k cf in let ndef = Array.length vt in - let cft = Array.map (fun v -> nf_val env v crazy_type) vt in + let cft = Array.map (fun v -> nf_val env sigma v crazy_type) vt in let name = Array.init ndef (fun _ -> (Name (Id.of_string "Fcofix"))) in let env = push_rec_types (name,cft,cft) env in - let cfb = Util.Array.map2 (fun v t -> nf_val env v t) vb cft in + let cfb = Util.Array.map2 (fun v t -> nf_val env sigma v t) vb cft in mkCoFix (init,(name,cft,cfb)) -let cbv_vm env c t = +let cbv_vm env sigma c t = let v = Vconv.val_of_constr env c in - nf_val env v t + nf_val env sigma v t let vm_infer_conv ?(pb=Reduction.CUMUL) env sigma t1 t2 = Reductionops.infer_conv_gen (fun pb ~l2r sigma ts -> Vconv.vm_conv_gen pb) diff --git a/pretyping/vnorm.mli b/pretyping/vnorm.mli index 58f5b14e1..bc6eec851 100644 --- a/pretyping/vnorm.mli +++ b/pretyping/vnorm.mli @@ -10,4 +10,4 @@ open Term open Environ (** {6 Reduction functions } *) -val cbv_vm : env -> constr -> types -> constr +val cbv_vm : env -> Evd.evar_map -> constr -> types -> constr diff --git a/printing/prettyp.ml b/printing/prettyp.ml index b19f8376c..ab0ce7e56 100644 --- a/printing/prettyp.ml +++ b/printing/prettyp.ml @@ -73,7 +73,7 @@ let print_ref reduce ref = let typ = Global.type_of_global_unsafe ref in let typ = if reduce then - let ctx,ccl = Reductionops.splay_prod_assum (Global.env()) Evd.empty typ + let ctx,ccl = Reductionops.splay_prod_assum (Global.env()) Evd.empty (EConstr.of_constr typ) in it_mkProd_or_LetIn ccl ctx else typ in let univs = Global.universes_of_global ref in @@ -594,7 +594,7 @@ let gallina_print_context with_values = prec let gallina_print_eval red_fun env sigma _ {uj_val=trm;uj_type=typ} = - let ntrm = red_fun env sigma trm in + let ntrm = red_fun env sigma (EConstr.of_constr trm) in (str " = " ++ gallina_print_typed_value_in_env env sigma (ntrm,typ)) (******************************************) diff --git a/proofs/clenv.ml b/proofs/clenv.ml index 34bc83097..d64cd9fff 100644 --- a/proofs/clenv.ml +++ b/proofs/clenv.ml @@ -72,7 +72,7 @@ let clenv_get_type_of ce c = Retyping.get_type_of (cl_env ce) (cl_sigma ce) c exception NotExtensibleClause let clenv_push_prod cl = - let typ = whd_all (cl_env cl) (cl_sigma cl) (clenv_type cl) in + let typ = whd_all (cl_env cl) (cl_sigma cl) (EConstr.of_constr (clenv_type cl)) in let rec clrec typ = match kind_of_term typ with | Cast (t,_,_) -> clrec t | Prod (na,t,u) -> @@ -264,7 +264,7 @@ let clenv_unify_meta_types ?(flags=default_unify_flags ()) clenv = let clenv_unique_resolver ?(flags=default_unify_flags ()) clenv gl = let concl = Goal.V82.concl clenv.evd (sig_it gl) in - if isMeta (fst (decompose_appvect (whd_nored clenv.evd clenv.templtyp.rebus))) then + if isMeta (fst (decompose_appvect (whd_nored clenv.evd (EConstr.of_constr clenv.templtyp.rebus)))) then clenv_unify CUMUL ~flags (clenv_type clenv) concl (clenv_unify_meta_types ~flags clenv) else @@ -482,7 +482,7 @@ let error_already_defined b = (str "Position " ++ int n ++ str" already defined.") let clenv_unify_binding_type clenv c t u = - if isMeta (fst (decompose_appvect (whd_nored clenv.evd u))) then + if isMeta (fst (decompose_appvect (whd_nored clenv.evd (EConstr.of_constr u)))) then (* Not enough information to know if some subtyping is needed *) CoerceToType, clenv, c else @@ -498,8 +498,8 @@ let clenv_unify_binding_type clenv c t u = TypeNotProcessed, clenv, c let clenv_assign_binding clenv k c = - let k_typ = clenv_hnf_constr clenv (clenv_meta_type clenv k) in - let c_typ = nf_betaiota clenv.evd (clenv_get_type_of clenv c) in + let k_typ = clenv_hnf_constr clenv (EConstr.of_constr (clenv_meta_type clenv k)) in + let c_typ = nf_betaiota clenv.evd (EConstr.of_constr (clenv_get_type_of clenv c)) in let status,clenv',c = clenv_unify_binding_type clenv c c_typ k_typ in { clenv' with evd = meta_assign k (c,(Conv,status)) clenv'.evd } diff --git a/proofs/logic.ml b/proofs/logic.ml index 29f295682..0fa193065 100644 --- a/proofs/logic.ml +++ b/proofs/logic.ml @@ -345,7 +345,7 @@ let rec mk_refgoals sigma goal goalacc conclty trm = else match kind_of_term trm with | Meta _ -> - let conclty = nf_betaiota sigma conclty in + let conclty = nf_betaiota sigma (EConstr.of_constr conclty) in if !check && occur_meta sigma (EConstr.of_constr conclty) then raise (RefinerError (MetaInType conclty)); let (gl,ev,sigma) = mk_goal hyps conclty in @@ -423,7 +423,7 @@ and mk_hdgoals sigma goal goalacc trm = match kind_of_term trm with | Cast (c,_, ty) when isMeta c -> check_typability env sigma ty; - let (gl,ev,sigma) = mk_goal hyps (nf_betaiota sigma ty) in + let (gl,ev,sigma) = mk_goal hyps (nf_betaiota sigma (EConstr.of_constr ty)) in gl::goalacc,ty,sigma,ev | Cast (t,_, ty) -> @@ -470,7 +470,7 @@ and mk_hdgoals sigma goal goalacc trm = and mk_arggoals sigma goal goalacc funty allargs = let foldmap (goalacc, funty, sigma) harg = - let t = whd_all (Goal.V82.env sigma goal) sigma funty in + let t = whd_all (Goal.V82.env sigma goal) sigma (EConstr.of_constr funty) in let rec collapse t = match kind_of_term t with | LetIn (_, c1, _, b) -> collapse (subst1 c1 b) | _ -> t @@ -491,21 +491,21 @@ and mk_casegoals sigma goal goalacc p c = let indspec = try Tacred.find_hnf_rectype env sigma ct with Not_found -> anomaly (Pp.str "mk_casegoals") in - let (lbrty,conclty) = type_case_branches_with_names env indspec p c in + let (lbrty,conclty) = type_case_branches_with_names env sigma indspec p c in (acc'',lbrty,conclty,sigma,p',c') let convert_hyp check sign sigma d = let id = NamedDecl.get_id d in - let b = NamedDecl.get_value d in + let b = Option.map EConstr.of_constr (NamedDecl.get_value d) in let env = Global.env() in let reorder = ref [] in let sign' = apply_to_hyp check sign id (fun _ d' _ -> - let c = NamedDecl.get_value d' in + let c = Option.map EConstr.of_constr (NamedDecl.get_value d') in let env = Global.env_of_context sign in - if check && not (is_conv env sigma (NamedDecl.get_type d) (NamedDecl.get_type d')) then + if check && not (is_conv env sigma (EConstr.of_constr (NamedDecl.get_type d)) (EConstr.of_constr (NamedDecl.get_type d'))) then user_err ~hdr:"Logic.convert_hyp" (str "Incorrect change of the type of " ++ pr_id id ++ str "."); if check && not (Option.equal (is_conv env sigma) b c) then @@ -532,7 +532,7 @@ let prim_refiner r sigma goal = (* Logical rules *) | Cut (b,replace,id,t) -> (* if !check && not (Retyping.get_sort_of env sigma t) then*) - let (sg1,ev1,sigma) = mk_goal sign (nf_betaiota sigma t) in + let (sg1,ev1,sigma) = mk_goal sign (nf_betaiota sigma (EConstr.of_constr t)) in let sign,t,cl,sigma = if replace then let nexthyp = get_hyp_after id (named_context_of_val sign) in diff --git a/proofs/redexpr.ml b/proofs/redexpr.ml index a442a5e63..40a8077a7 100644 --- a/proofs/redexpr.ml +++ b/proofs/redexpr.ml @@ -24,10 +24,11 @@ open Misctypes (* call by value normalisation function using the virtual machine *) let cbv_vm env sigma c = + let c = EConstr.Unsafe.to_constr c in let ctyp = Retyping.get_type_of env sigma c in if Termops.occur_meta_or_existential sigma (EConstr.of_constr c) then error "vm_compute does not support existential variables."; - Vnorm.cbv_vm env c ctyp + Vnorm.cbv_vm env sigma c ctyp let warn_native_compute_disabled = CWarnings.create ~name:"native-compute-disabled" ~category:"native-compiler" @@ -39,13 +40,15 @@ let cbv_native env sigma c = (warn_native_compute_disabled (); cbv_vm env sigma c) else + let c = EConstr.Unsafe.to_constr c in let ctyp = Retyping.get_type_of env sigma c in Nativenorm.native_norm env sigma c ctyp let whd_cbn flags env sigma t = let (state,_) = (whd_state_gen true true flags env sigma (t,Reductionops.Stack.empty)) - in Reductionops.Stack.zip ~refold:true state + in + EConstr.Unsafe.to_constr (Reductionops.Stack.zip ~refold:true sigma state) let strong_cbn flags = strong (whd_cbn flags) diff --git a/proofs/tacmach.ml b/proofs/tacmach.ml index e41fb94cc..09f5246ab 100644 --- a/proofs/tacmach.ml +++ b/proofs/tacmach.ml @@ -77,26 +77,27 @@ let pf_global gls id = Constrintern.construct_reference (pf_hyps gls) id let pf_reduction_of_red_expr gls re c = let (redfun, _) = reduction_of_red_expr (pf_env gls) re in let sigma = Sigma.Unsafe.of_evar_map (project gls) in - let Sigma (c, sigma, _) = redfun.e_redfun (pf_env gls) sigma c in + let Sigma (c, sigma, _) = redfun.e_redfun (pf_env gls) sigma (EConstr.of_constr c) in (Sigma.to_evar_map sigma, c) let pf_apply f gls = f (pf_env gls) (project gls) let pf_eapply f gls x = on_sig gls (fun evm -> f (pf_env gls) evm x) let pf_reduce = pf_apply +let pf_reduce' f gl c = pf_apply f gl (EConstr.of_constr c) let pf_e_reduce = pf_apply -let pf_whd_all = pf_reduce whd_all -let pf_hnf_constr = pf_reduce hnf_constr -let pf_nf = pf_reduce simpl -let pf_nf_betaiota = pf_reduce (fun _ -> nf_betaiota) -let pf_compute = pf_reduce compute -let pf_unfoldn ubinds = pf_reduce (unfoldn ubinds) +let pf_whd_all = pf_reduce' whd_all +let pf_hnf_constr = pf_reduce' hnf_constr +let pf_nf = pf_reduce' simpl +let pf_nf_betaiota = pf_reduce' (fun _ -> nf_betaiota) +let pf_compute = pf_reduce' compute +let pf_unfoldn ubinds = pf_reduce' (unfoldn ubinds) let pf_unsafe_type_of = pf_reduce unsafe_type_of let pf_type_of = pf_reduce type_of let pf_get_type_of = pf_reduce Retyping.get_type_of -let pf_conv_x gl = pf_reduce test_conversion gl Reduction.CONV +let pf_conv_x gl = pf_apply test_conversion gl Reduction.CONV let pf_conv_x_leq gl = pf_reduce test_conversion gl Reduction.CUMUL let pf_const_value = pf_reduce (fun env _ -> constant_value_in env) @@ -158,6 +159,9 @@ module New = struct let pf_apply f gl = f (Proofview.Goal.env gl) (project gl) + let pf_reduce f gl c = + f (Proofview.Goal.env gl) (project gl) (EConstr.of_constr c) + let of_old f gl = f { Evd.it = Proofview.Goal.goal gl ; sigma = project gl; } @@ -217,14 +221,14 @@ module New = struct let sigma = project gl in nf_evar sigma concl - let pf_whd_all gl t = pf_apply whd_all gl t + let pf_whd_all gl t = pf_reduce whd_all gl t let pf_get_type_of gl t = pf_apply Retyping.get_type_of gl t let pf_reduce_to_quantified_ind gl t = pf_apply reduce_to_quantified_ind gl t - let pf_hnf_constr gl t = pf_apply hnf_constr gl t + let pf_hnf_constr gl t = pf_reduce hnf_constr gl t let pf_hnf_type_of gl t = pf_whd_all gl (pf_get_type_of gl t) diff --git a/proofs/tacmach.mli b/proofs/tacmach.mli index 727efcf6d..16f6f88c1 100644 --- a/proofs/tacmach.mli +++ b/proofs/tacmach.mli @@ -110,7 +110,7 @@ module New : sig val pf_unsafe_type_of : ('a, 'r) Proofview.Goal.t -> Term.constr -> Term.types val pf_type_of : ('a, 'r) Proofview.Goal.t -> Term.constr -> evar_map * Term.types - val pf_conv_x : ('a, 'r) Proofview.Goal.t -> Term.constr -> Term.constr -> bool + val pf_conv_x : ('a, 'r) Proofview.Goal.t -> EConstr.t -> EConstr.t -> bool val pf_get_new_id : identifier -> ([ `NF ], 'r) Proofview.Goal.t -> identifier val pf_ids_of_hyps : ('a, 'r) Proofview.Goal.t -> identifier list @@ -126,8 +126,8 @@ module New : sig val pf_hnf_constr : ('a, 'r) Proofview.Goal.t -> constr -> types val pf_hnf_type_of : ('a, 'r) Proofview.Goal.t -> constr -> types - val pf_whd_all : ('a, 'r) Proofview.Goal.t -> constr -> constr - val pf_compute : ('a, 'r) Proofview.Goal.t -> constr -> constr + val pf_whd_all : ('a, 'r) Proofview.Goal.t -> EConstr.t -> constr + val pf_compute : ('a, 'r) Proofview.Goal.t -> EConstr.t -> constr val pf_matches : ('a, 'r) Proofview.Goal.t -> constr_pattern -> constr -> patvar_map diff --git a/stm/lemmas.ml b/stm/lemmas.ml index 55f33be39..9896d5a93 100644 --- a/stm/lemmas.ml +++ b/stm/lemmas.ml @@ -108,7 +108,7 @@ let find_mutually_recursive_statements thms = (* Cofixpoint or fixpoint w/o explicit decreasing argument *) | None | Some (None, CStructRec) -> let whnf_hyp_hds = map_rel_context_in_env - (fun env c -> fst (whd_all_stack env Evd.empty c)) + (fun env c -> EConstr.Unsafe.to_constr (fst (whd_all_stack env Evd.empty (EConstr.of_constr c)))) (Global.env()) hyps in let ind_hyps = List.flatten (List.map_i (fun i decl -> @@ -122,8 +122,8 @@ let find_mutually_recursive_statements thms = []) 0 (List.rev whnf_hyp_hds)) in let ind_ccl = let cclenv = push_rel_context hyps (Global.env()) in - let whnf_ccl,_ = whd_all_stack cclenv Evd.empty ccl in - match kind_of_term whnf_ccl with + let whnf_ccl,_ = whd_all_stack cclenv Evd.empty (EConstr.of_constr ccl) in + match kind_of_term (EConstr.Unsafe.to_constr whnf_ccl) with | Ind ((kn,_ as ind),u) when let mind = Global.lookup_mind kn in Int.equal mind.mind_ntypes n && mind.mind_finite == Decl_kinds.CoFinite -> diff --git a/tactics/autorewrite.ml b/tactics/autorewrite.ml index dae1cc9f1..46600cdd7 100644 --- a/tactics/autorewrite.ml +++ b/tactics/autorewrite.ml @@ -284,7 +284,7 @@ let decompose_applied_relation metas env sigma c ctype left2right = match find_rel ctype with | Some c -> Some c | None -> - let ctx,t' = Reductionops.splay_prod_assum env sigma ctype in (* Search for underlying eq *) + let ctx,t' = Reductionops.splay_prod_assum env sigma (EConstr.of_constr ctype) in (* Search for underlying eq *) match find_rel (it_mkProd_or_LetIn t' ctx) with | Some c -> Some c | None -> None diff --git a/tactics/contradiction.ml b/tactics/contradiction.ml index fcbad4bf0..b9704b846 100644 --- a/tactics/contradiction.ml +++ b/tactics/contradiction.ml @@ -65,7 +65,7 @@ let contradiction_context = | d :: rest -> let id = NamedDecl.get_id d in let typ = nf_evar sigma (NamedDecl.get_type d) in - let typ = whd_all env sigma typ in + let typ = whd_all env sigma (EConstr.of_constr typ) in if is_empty_type sigma typ then simplest_elim (mkVar id) else match kind_of_term typ with @@ -88,7 +88,7 @@ let contradiction_context = (Proofview.tclORELSE (Proofview.Goal.enter { enter = begin fun gl -> let is_conv_leq = Tacmach.New.pf_apply is_conv_leq gl in - filter_hyp (fun typ -> is_conv_leq typ t) + filter_hyp (fun typ -> is_conv_leq (EConstr.of_constr typ) (EConstr.of_constr t)) (fun id' -> simplest_elim (mkApp (mkVar id,[|mkVar id'|]))) end }) begin function (e, info) -> match e with @@ -105,7 +105,7 @@ let is_negation_of env sigma typ t = match kind_of_term (whd_all env sigma t) with | Prod (na,t,u) -> let u = nf_evar sigma u in - is_empty_type sigma u && is_conv_leq env sigma typ t + is_empty_type sigma u && is_conv_leq env sigma (EConstr.of_constr typ) (EConstr.of_constr t) | _ -> false let contradiction_term (c,lbind as cl) = @@ -114,7 +114,7 @@ let contradiction_term (c,lbind as cl) = let env = Proofview.Goal.env gl in let type_of = Tacmach.New.pf_unsafe_type_of gl in let typ = type_of c in - let _, ccl = splay_prod env sigma typ in + let _, ccl = splay_prod env sigma (EConstr.of_constr typ) in if is_empty_type sigma ccl then Tacticals.New.tclTHEN (elim false None cl None) @@ -123,7 +123,7 @@ let contradiction_term (c,lbind as cl) = Proofview.tclORELSE begin if lbind = NoBindings then - filter_hyp (is_negation_of env sigma typ) + filter_hyp (fun c -> is_negation_of env sigma typ (EConstr.of_constr c)) (fun id -> simplest_elim (mkApp (mkVar id,[|c|]))) else Proofview.tclZERO Not_found diff --git a/tactics/eauto.ml b/tactics/eauto.ml index 6250fef2d..0869ac0c7 100644 --- a/tactics/eauto.ml +++ b/tactics/eauto.ml @@ -478,7 +478,7 @@ let unfold_head env (ids, csts) c = true, Environ.constant_value_in env c | App (f, args) -> (match aux f with - | true, f' -> true, Reductionops.whd_betaiota Evd.empty (mkApp (f', args)) + | true, f' -> true, Reductionops.whd_betaiota Evd.empty (EConstr.of_constr (mkApp (f', args))) | false, _ -> let done_, args' = Array.fold_left_i (fun i (done_, acc) arg -> diff --git a/tactics/elim.ml b/tactics/elim.ml index 12d8e98c4..b830ccefe 100644 --- a/tactics/elim.ml +++ b/tactics/elim.ml @@ -96,14 +96,14 @@ let head_in indl t gl = let ity,_ = if !up_to_delta then find_mrectype env sigma t - else extract_mrectype t + else extract_mrectype sigma t in List.exists (fun i -> eq_ind (fst i) (fst ity)) indl with Not_found -> false let decompose_these c l = Proofview.Goal.enter { enter = begin fun gl -> let indl = List.map (fun x -> x, Univ.Instance.empty) l in - general_decompose (fun sigma (_,t) -> head_in indl t gl) c + general_decompose (fun sigma (_,t) -> head_in indl (EConstr.of_constr t) gl) c end } let decompose_and c = diff --git a/tactics/eqdecide.ml b/tactics/eqdecide.ml index 1a67bedc2..1554d43f0 100644 --- a/tactics/eqdecide.ml +++ b/tactics/eqdecide.ml @@ -197,7 +197,7 @@ let decideGralEquality = Proofview.Goal.enter { enter = begin fun gl -> let concl = pf_nf_concl gl in match_eqdec concl >>= fun (eqonleft,_,c1,c2,typ) -> - let headtyp = hd_app (pf_compute gl typ) in + let headtyp = hd_app (pf_compute gl (EConstr.of_constr typ)) in begin match kind_of_term headtyp with | Ind (mi,_) -> Proofview.tclUNIT mi | _ -> tclZEROMSG (Pp.str"This decision procedure only works for inductive objects.") diff --git a/tactics/eqschemes.ml b/tactics/eqschemes.ml index c94dcfa9d..aea3ca17e 100644 --- a/tactics/eqschemes.ml +++ b/tactics/eqschemes.ml @@ -606,8 +606,8 @@ let fix_r2l_forward_rew_scheme (c, ctx') = (mkLambda_or_LetIn (RelDecl.map_constr (liftn (-1) 2) hp) (mkLambda_or_LetIn (RelDecl.map_constr (lift 2) ind) (Reductionops.whd_beta Evd.empty - (applist (c, - Context.Rel.to_extended_list 3 indargs @ [mkRel 1;mkRel 3;mkRel 2])))))) + (EConstr.of_constr (applist (c, + Context.Rel.to_extended_list 3 indargs @ [mkRel 1;mkRel 3;mkRel 2]))))))) in c', ctx' | _ -> anomaly (Pp.str "Ill-formed non-dependent left-to-right rewriting scheme") diff --git a/tactics/equality.ml b/tactics/equality.ml index 74f6dd44a..48f46b36b 100644 --- a/tactics/equality.ml +++ b/tactics/equality.ml @@ -328,7 +328,7 @@ let jmeq_same_dom gl = function | Some t -> let rels, t = decompose_prod_assum t in let env = Environ.push_rel_context rels (Proofview.Goal.env gl) in - match decompose_app t with + match EConstr.decompose_app (project gl) (EConstr.of_constr t) with | _, [dom1; _; dom2;_] -> is_conv env (Tacmach.New.project gl) dom1 dom2 | _ -> false @@ -402,7 +402,7 @@ let type_of_clause cls gl = match cls with let leibniz_rewrite_ebindings_clause cls lft2rgt tac c t l with_evars frzevars dep_proof_ok hdcncl = Proofview.Goal.nf_s_enter { s_enter = begin fun gl -> let evd = Sigma.to_evar_map (Proofview.Goal.sigma gl) in - let isatomic = isProd (whd_zeta evd hdcncl) in + let isatomic = isProd (whd_zeta evd (EConstr.of_constr hdcncl)) in let dep_fun = if isatomic then dependent else dependent_no_evar in let type_of_cls = type_of_clause cls gl in let dep = dep_proof_ok && dep_fun evd (EConstr.of_constr c) (EConstr.of_constr type_of_cls) in @@ -441,7 +441,7 @@ let general_rewrite_ebindings_clause cls lft2rgt occs frzevars dep_proof_ok ?tac let sigma = Tacmach.New.project gl in let env = Proofview.Goal.env gl in let ctype = get_type_of env sigma c in - let rels, t = decompose_prod_assum (whd_betaiotazeta sigma ctype) in + let rels, t = decompose_prod_assum (whd_betaiotazeta sigma (EConstr.of_constr ctype)) in match match_with_equality_type sigma t with | Some (hdcncl,args) -> (* Fast path: direct leibniz-like rewrite *) let lft2rgt = adjust_rewriting_direction args lft2rgt in @@ -457,7 +457,7 @@ let general_rewrite_ebindings_clause cls lft2rgt occs frzevars dep_proof_ok ?tac | (e, info) -> Proofview.tclEVARMAP >>= fun sigma -> let env' = push_rel_context rels env in - let rels',t' = splay_prod_assum env' sigma t in (* Search for underlying eq *) + let rels',t' = splay_prod_assum env' sigma (EConstr.of_constr t) in (* Search for underlying eq *) match match_with_equality_type sigma t' with | Some (hdcncl,args) -> let lft2rgt = adjust_rewriting_direction args lft2rgt in @@ -714,9 +714,11 @@ let _ = optread = (fun () -> !keep_proof_equalities_for_injection) ; optwrite = (fun b -> keep_proof_equalities_for_injection := b) } - let find_positions env sigma t1 t2 = + let open EConstr in let project env sorts posn t1 t2 = + let t1 = EConstr.Unsafe.to_constr t1 in + let t2 = EConstr.Unsafe.to_constr t2 in let ty1 = get_type_of env sigma t1 in let s = get_sort_family_of env sigma ty1 in if Sorts.List.mem s sorts @@ -725,7 +727,7 @@ let find_positions env sigma t1 t2 = let rec findrec sorts posn t1 t2 = let hd1,args1 = whd_all_stack env sigma t1 in let hd2,args2 = whd_all_stack env sigma t2 in - match (kind_of_term hd1, kind_of_term hd2) with + match (EConstr.kind sigma hd1, EConstr.kind sigma hd2) with | Construct (sp1,_), Construct (sp2,_) when Int.equal (List.length args1) (constructor_nallargs_env env sp1) -> @@ -760,7 +762,7 @@ let find_positions env sigma t1 t2 = let sorts = if !keep_proof_equalities_for_injection then [InSet;InType;InProp] else [InSet;InType] in - Inr (findrec sorts [] t1 t2) + Inr (findrec sorts [] (EConstr.of_constr t1) (EConstr.of_constr t2)) with DiscrFound (path,c1,c2) -> Inl (path,c1,c2) @@ -840,7 +842,7 @@ let injectable env sigma t1 t2 = let descend_then env sigma head dirn = let IndType (indf,_) = - try find_rectype env sigma (get_type_of env sigma head) + try find_rectype env sigma (EConstr.of_constr (get_type_of env sigma head)) with Not_found -> error "Cannot project on an inductive type derived from a dependency." in let indp,_ = (dest_ind_family indf) in @@ -883,7 +885,7 @@ let descend_then env sigma head dirn = let build_selector env sigma dirn c ind special default = let IndType(indf,_) = - try find_rectype env sigma ind + try find_rectype env sigma (EConstr.of_constr ind) with Not_found -> (* one can find Rel(k) in case of dependent constructors like T := c : (A:Set)A->T and a discrimination @@ -1026,7 +1028,7 @@ let onNegatedEquality with_evars tac = let sigma = Tacmach.New.project gl in let ccl = Proofview.Goal.concl gl in let env = Proofview.Goal.env gl in - match kind_of_term (hnf_constr env sigma ccl) with + match kind_of_term (hnf_constr env sigma (EConstr.of_constr ccl)) with | Prod (_,t,u) when is_empty_type sigma u -> tclTHEN introf (onLastHypId (fun id -> @@ -1104,7 +1106,7 @@ let make_tuple env sigma (rterm,rty) lind = let minimal_free_rels env sigma (c,cty) = let cty_rels = free_rels sigma (EConstr.of_constr cty) in - let cty' = simpl env sigma cty in + let cty' = simpl env sigma (EConstr.of_constr cty) in let rels' = free_rels sigma (EConstr.of_constr cty') in if Int.Set.subset cty_rels rels' then (cty,cty_rels) @@ -1171,11 +1173,11 @@ let sig_clausal_form env sigma sort_of_ty siglen ty dflt = with Evarconv.UnableToUnify _ -> error "Cannot solve a unification problem." else - let (a,p_i_minus_1) = match whd_beta_stack !evdref p_i with - | (_sigS,[a;p]) -> (a,p) + let (a,p_i_minus_1) = match whd_beta_stack !evdref (EConstr.of_constr p_i) with + | (_sigS,[a;p]) -> (EConstr.Unsafe.to_constr a, EConstr.Unsafe.to_constr p) | _ -> anomaly ~label:"sig_clausal_form" (Pp.str "should be a sigma type") in let ev = Evarutil.e_new_evar env evdref a in - let rty = beta_applist(p_i_minus_1,[ev]) in + let rty = beta_applist sigma (EConstr.of_constr p_i_minus_1,[EConstr.of_constr ev]) in let tuple_tail = sigrec_clausal_form (siglen-1) rty in match Evd.existential_opt_value !evdref @@ -1317,13 +1319,13 @@ let inject_if_homogenous_dependent_pair ty = hd2,ar2 = decompose_app_vect sigma (EConstr.of_constr t2) in if not (Globnames.is_global (existTconstr()) hd1) then raise Exit; if not (Globnames.is_global (existTconstr()) hd2) then raise Exit; - let ind,_ = try pf_apply find_mrectype gl ar1.(0) with Not_found -> raise Exit in + let ind,_ = try pf_apply find_mrectype gl (EConstr.of_constr ar1.(0)) with Not_found -> raise Exit in (* check if the user has declared the dec principle *) (* and compare the fst arguments of the dep pair *) (* Note: should work even if not an inductive type, but the table only *) (* knows inductive types *) if not (Ind_tables.check_scheme (!eq_dec_scheme_kind_name()) (fst ind) && - pf_apply is_conv gl ar1.(2) ar2.(2)) then raise Exit; + pf_apply is_conv gl (EConstr.of_constr ar1.(2)) (EConstr.of_constr ar2.(2))) then raise Exit; Coqlib.check_required_library ["Coq";"Logic";"Eqdep_dec"]; let new_eq_args = [|pf_unsafe_type_of gl ar1.(3);ar1.(3);ar2.(3)|] in let inj2 = Coqlib.coq_constant "inj_pair2_eq_dec is missing" @@ -1350,8 +1352,8 @@ let inject_if_homogenous_dependent_pair ty = let simplify_args env sigma t = (* Quick hack to reduce in arguments of eq only *) match decompose_app t with - | eq, [t;c1;c2] -> applist (eq,[t;simpl env sigma c1;simpl env sigma c2]) - | eq, [t1;c1;t2;c2] -> applist (eq,[t1;simpl env sigma c1;t2;simpl env sigma c2]) + | eq, [t;c1;c2] -> applist (eq,[t;simpl env sigma (EConstr.of_constr c1);simpl env sigma (EConstr.of_constr c2)]) + | eq, [t1;c1;t2;c2] -> applist (eq,[t1;simpl env sigma (EConstr.of_constr c1);t2;simpl env sigma (EConstr.of_constr c2)]) | _ -> t let inject_at_positions env sigma l2r (eq,_,(t,t1,t2)) eq_clause posns tac = @@ -1515,14 +1517,14 @@ let _ = declare_intro_decomp_eq intro_decomp_eq *) -let decomp_tuple_term env c t = +let decomp_tuple_term env sigma c t = let rec decomprec inner_code ex exty = let iterated_decomp = try let ({proj1=p1; proj2=p2}),(i,a,p,car,cdr) = find_sigma_data_decompose ex in let car_code = applist (mkConstU (destConstRef p1,i),[a;p;inner_code]) and cdr_code = applist (mkConstU (destConstRef p2,i),[a;p;inner_code]) in - let cdrtyp = beta_applist (p,[car]) in + let cdrtyp = beta_applist sigma (EConstr.of_constr p,[EConstr.of_constr car]) in List.map (fun l -> ((car,a),car_code)::l) (decomprec cdr_code cdr cdrtyp) with Constr_matching.PatternMatchingFailure -> [] @@ -1533,8 +1535,8 @@ let subst_tuple_term env sigma dep_pair1 dep_pair2 b = let sigma = Sigma.to_evar_map sigma in let typ = get_type_of env sigma dep_pair1 in (* We find all possible decompositions *) - let decomps1 = decomp_tuple_term env dep_pair1 typ in - let decomps2 = decomp_tuple_term env dep_pair2 typ in + let decomps1 = decomp_tuple_term env sigma dep_pair1 typ in + let decomps2 = decomp_tuple_term env sigma dep_pair2 typ in (* We adjust to the shortest decomposition *) let n = min (List.length decomps1) (List.length decomps2) in let decomp1 = List.nth decomps1 (n-1) in @@ -1547,11 +1549,11 @@ let subst_tuple_term env sigma dep_pair1 dep_pair2 b = let abst_B = List.fold_right (fun (e,t) body -> lambda_create env (t,subst_term sigma (EConstr.of_constr e) (EConstr.of_constr body))) e1_list b in - let pred_body = beta_applist(abst_B,proj_list) in + let pred_body = beta_applist sigma (EConstr.of_constr abst_B, List.map EConstr.of_constr proj_list) in let body = mkApp (lambda_create env (typ,pred_body),[|dep_pair1|]) in - let expected_goal = beta_applist (abst_B,List.map fst e2_list) in + let expected_goal = beta_applist sigma (EConstr.of_constr abst_B,List.map (fst %> EConstr.of_constr) e2_list) in (* Simulate now the normalisation treatment made by Logic.mk_refgoals *) - let expected_goal = nf_betaiota sigma expected_goal in + let expected_goal = nf_betaiota sigma (EConstr.of_constr expected_goal) in (* Retype to get universes right *) let sigma, expected_goal_ty = Typing.type_of env sigma expected_goal in let sigma, _ = Typing.type_of env sigma body in @@ -1842,20 +1844,20 @@ let subst_all ?(flags=default_subst_tactic_flags ()) () = let cond_eq_term_left c t gl = try let (_,x,_) = pi3 (find_eq_data_decompose gl t) in - if pf_conv_x gl c x then true else failwith "not convertible" + if pf_conv_x gl (EConstr.of_constr c) (EConstr.of_constr x) then true else failwith "not convertible" with Constr_matching.PatternMatchingFailure -> failwith "not an equality" let cond_eq_term_right c t gl = try let (_,_,x) = pi3 (find_eq_data_decompose gl t) in - if pf_conv_x gl c x then false else failwith "not convertible" + if pf_conv_x gl (EConstr.of_constr c) (EConstr.of_constr x) then false else failwith "not convertible" with Constr_matching.PatternMatchingFailure -> failwith "not an equality" let cond_eq_term c t gl = try let (_,x,y) = pi3 (find_eq_data_decompose gl t) in - if pf_conv_x gl c x then true - else if pf_conv_x gl c y then false + if pf_conv_x gl (EConstr.of_constr c) (EConstr.of_constr x) then true + else if pf_conv_x gl (EConstr.of_constr c) (EConstr.of_constr y) then false else failwith "not convertible" with Constr_matching.PatternMatchingFailure -> failwith "not an equality" diff --git a/tactics/hints.ml b/tactics/hints.ml index 55bf5f29e..c41f88ab7 100644 --- a/tactics/hints.ml +++ b/tactics/hints.ml @@ -761,7 +761,7 @@ let make_exact_entry env sigma pri poly ?(name=PathAny) (c, cty, ctx) = code = with_uid (Give_exact (c, cty, ctx)); }) let make_apply_entry env sigma (eapply,hnf,verbose) pri poly ?(name=PathAny) (c, cty, ctx) = - let cty = if hnf then hnf_constr env sigma cty else cty in + let cty = if hnf then hnf_constr env sigma (EConstr.of_constr cty) else cty in match kind_of_term cty with | Prod _ -> let sigma' = Evd.merge_context_set univ_flexible sigma ctx in @@ -910,7 +910,7 @@ let make_mode ref m = let make_trivial env sigma poly ?(name=PathAny) r = let c,ctx = fresh_global_or_constr env sigma poly r in let sigma = Evd.merge_context_set univ_flexible sigma ctx in - let t = hnf_constr env sigma (unsafe_type_of env sigma c) in + let t = hnf_constr env sigma (EConstr.of_constr (unsafe_type_of env sigma c)) in let hd = head_of_constr_reference (head_constr sigma t) in let ce = mk_clenv_from_env env sigma None (c,t) in (Some hd, { pri=1; diff --git a/tactics/hipattern.ml b/tactics/hipattern.ml index 847ecf4b0..a42a51fc0 100644 --- a/tactics/hipattern.ml +++ b/tactics/hipattern.ml @@ -440,7 +440,7 @@ let extract_eq_args gl = function let t = pf_unsafe_type_of gl e1 in (t,e1,e2) | PolymorphicLeibnizEq (t,e1,e2) -> (t,e1,e2) | HeterogenousEq (t1,e1,t2,e2) -> - if pf_conv_x gl t1 t2 then (t1,e1,e2) + if pf_conv_x gl (EConstr.of_constr t1) (EConstr.of_constr t2) then (t1,e1,e2) else raise PatternMatchingFailure let find_eq_data_decompose gl eqn = @@ -466,7 +466,7 @@ let match_eq_nf gls eqn (ref, hetero) = match Id.Map.bindings (pf_matches gls pat eqn) with | [(m1,t);(m2,x);(m3,y)] -> assert (Id.equal m1 meta1 && Id.equal m2 meta2 && Id.equal m3 meta3); - (t,pf_whd_all gls x,pf_whd_all gls y) + (t,pf_whd_all gls (EConstr.of_constr x),pf_whd_all gls (EConstr.of_constr y)) | _ -> anomaly ~label:"match_eq" (Pp.str "an eq pattern should match 3 terms") let dest_nf_eq gls eqn = diff --git a/tactics/inv.ml b/tactics/inv.ml index d1d6178da..0b2d2f0b2 100644 --- a/tactics/inv.ml +++ b/tactics/inv.ml @@ -443,7 +443,7 @@ let raw_inversion inv_kind id status names = let msg = str "The type of " ++ pr_id id ++ str " is not inductive." in CErrors.user_err msg in - let IndType (indf,realargs) = find_rectype env sigma t in + let IndType (indf,realargs) = find_rectype env sigma (EConstr.of_constr t) in let evdref = ref sigma in let (elim_predicate, args) = make_inv_predicate env evdref indf realargs id status concl in diff --git a/tactics/leminv.ml b/tactics/leminv.ml index 46f1f7c8d..85910355e 100644 --- a/tactics/leminv.ml +++ b/tactics/leminv.ml @@ -116,7 +116,7 @@ let max_prefix_sign lid sign = | id::l -> snd (max_rec (id, sign_prefix id sign) l) *) let rec add_prods_sign env sigma t = - match kind_of_term (whd_all env sigma t) with + match kind_of_term (whd_all env sigma (EConstr.of_constr t)) with | Prod (na,c1,b) -> let id = id_of_name_using_hdchar env t na in let b'= subst1 (mkVar id) b in @@ -169,7 +169,7 @@ let compute_first_inversion_scheme env sigma ind sort dep_option = let goal = mkArrow i (applist(mkVar p, List.rev revargs)) in (pty,goal) in - let npty = nf_all env sigma pty in + let npty = nf_all env sigma (EConstr.of_constr pty) in let extenv = push_named (LocalAssum (p,npty)) env in extenv, goal @@ -183,7 +183,7 @@ let compute_first_inversion_scheme env sigma ind sort dep_option = let inversion_scheme env sigma t sort dep_option inv_op = let (env,i) = add_prods_sign env sigma t in let ind = - try find_rectype env sigma i + try find_rectype env sigma (EConstr.of_constr i) with Not_found -> user_err ~hdr:"inversion_scheme" (no_inductive_inconstr env sigma i) in diff --git a/tactics/tactics.ml b/tactics/tactics.ml index 15dd1a97c..c96553fae 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -502,7 +502,7 @@ let rec check_mutind env sigma k cl = match kind_of_term (strip_outer_cast sigma | Prod (na, c1, b) -> if Int.equal k 1 then try - let ((sp, _), u), _ = find_inductive env sigma c1 in + let ((sp, _), u), _ = find_inductive env sigma (EConstr.of_constr c1) in (sp, u) with Not_found -> error "Cannot do a fixpoint on a non inductive type." else @@ -555,14 +555,14 @@ let fix ido n = match ido with mutual_fix id n [] 0 let rec check_is_mutcoind env sigma cl = - let b = whd_all env sigma cl in + let b = whd_all env sigma (EConstr.of_constr cl) in match kind_of_term b with | Prod (na, c1, b) -> let open Context.Rel.Declaration in check_is_mutcoind (push_rel (LocalAssum (na,c1)) env) sigma b | _ -> try - let _ = find_coinductive env sigma b in () + let _ = find_coinductive env sigma (EConstr.of_constr b) in () with Not_found -> error "All methods must construct elements in coinductive types." @@ -609,11 +609,11 @@ let cofix ido = match ido with (* Reduction and conversion tactics *) (**************************************************************) -type tactic_reduction = env -> evar_map -> constr -> constr +type tactic_reduction = env -> evar_map -> EConstr.t -> constr let pf_reduce_decl redfun where decl gl = let open Context.Named.Declaration in - let redfun' = Tacmach.New.pf_apply redfun gl in + let redfun' c = Tacmach.New.pf_apply redfun gl (EConstr.of_constr c) in match decl with | LocalAssum (id,ty) -> if where == InHypValueOnly then @@ -694,7 +694,7 @@ let bind_red_expr_occurrences occs nbcl redexp = let reduct_in_concl (redfun,sty) = Proofview.Goal.nf_enter { enter = begin fun gl -> - convert_concl_no_check (Tacmach.New.pf_apply redfun gl (Tacmach.New.pf_concl gl)) sty + convert_concl_no_check (Tacmach.New.pf_apply redfun gl (EConstr.of_constr (Tacmach.New.pf_concl gl))) sty end } let reduct_in_hyp ?(check=false) redfun (id,where) = @@ -714,7 +714,7 @@ let reduct_option ?(check=false) redfun = function let pf_e_reduce_decl redfun where decl gl = let open Context.Named.Declaration in let sigma = Proofview.Goal.sigma gl in - let redfun sigma c = redfun.e_redfun (Tacmach.New.pf_env gl) sigma c in + let redfun sigma c = redfun.e_redfun (Tacmach.New.pf_env gl) sigma (EConstr.of_constr c) in match decl with | LocalAssum (id,ty) -> if where == InHypValueOnly then @@ -729,7 +729,7 @@ let pf_e_reduce_decl redfun where decl gl = let e_reduct_in_concl ~check (redfun, sty) = Proofview.Goal.nf_s_enter { s_enter = begin fun gl -> let sigma = Proofview.Goal.sigma gl in - let Sigma (c', sigma, p) = redfun.e_redfun (Tacmach.New.pf_env gl) sigma (Tacmach.New.pf_concl gl) in + let Sigma (c', sigma, p) = redfun.e_redfun (Tacmach.New.pf_env gl) sigma (EConstr.of_constr (Tacmach.New.pf_concl gl)) in Sigma (convert_concl ~check c' sty, sigma, p) end } @@ -749,7 +749,7 @@ let e_reduct_option ?(check=false) redfun = function let e_change_in_concl (redfun,sty) = Proofview.Goal.s_enter { s_enter = begin fun gl -> let sigma = Proofview.Goal.sigma gl in - let Sigma (c, sigma, p) = redfun.e_redfun (Proofview.Goal.env gl) sigma (Proofview.Goal.raw_concl gl) in + let Sigma (c, sigma, p) = redfun.e_redfun (Proofview.Goal.env gl) sigma (EConstr.of_constr (Proofview.Goal.raw_concl gl)) in Sigma (convert_concl_no_check c sty, sigma, p) end } @@ -759,14 +759,14 @@ let e_pf_change_decl (redfun : bool -> e_reduction_function) where decl env sigm | LocalAssum (id,ty) -> if where == InHypValueOnly then user_err (pr_id id ++ str " has no value."); - let Sigma (ty', sigma, p) = (redfun false).e_redfun env sigma ty in + let Sigma (ty', sigma, p) = (redfun false).e_redfun env sigma (EConstr.of_constr ty) in Sigma (LocalAssum (id, ty'), sigma, p) | LocalDef (id,b,ty) -> let Sigma (b', sigma, p) = - if where != InHypTypeOnly then (redfun true).e_redfun env sigma b else Sigma.here b sigma + if where != InHypTypeOnly then (redfun true).e_redfun env sigma (EConstr.of_constr b) else Sigma.here b sigma in let Sigma (ty', sigma, q) = - if where != InHypValueOnly then (redfun false).e_redfun env sigma ty else Sigma.here ty sigma + if where != InHypValueOnly then (redfun false).e_redfun env sigma (EConstr.of_constr ty) else Sigma.here ty sigma in Sigma (LocalDef (id,b',ty'), sigma, p +> q) @@ -792,20 +792,21 @@ let check_types env sigma mayneedglobalcheck deep newc origc = let sigma, b = infer_conv ~pb:Reduction.CUMUL env sigma t1 t2 in if not b then if - isSort (whd_all env sigma t1) && - isSort (whd_all env sigma t2) + isSort (whd_all env sigma (EConstr.of_constr t1)) && + isSort (whd_all env sigma (EConstr.of_constr t2)) then (mayneedglobalcheck := true; sigma) else user_err ~hdr:"convert-check-hyp" (str "Types are incompatible.") else sigma end else - if not (isSort (whd_all env sigma t1)) then + if not (isSort (whd_all env sigma (EConstr.of_constr t1))) then user_err ~hdr:"convert-check-hyp" (str "Not a type.") else sigma (* Now we introduce different instances of the previous tacticals *) let change_and_check cv_pb mayneedglobalcheck deep t = { e_redfun = begin fun env sigma c -> + let c = EConstr.Unsafe.to_constr c in let Sigma (t', sigma, p) = t.run sigma in let sigma = Sigma.to_evar_map sigma in let sigma = check_types env sigma mayneedglobalcheck deep t' c in @@ -1079,7 +1080,7 @@ let lookup_hypothesis_as_renamed_gen red h gl = match lookup_hypothesis_as_renamed env ccl h with | None when red -> let (redfun, _) = Redexpr.reduction_of_red_expr env (Red true) in - let Sigma (c, _, _) = redfun.e_redfun env (Proofview.Goal.sigma gl) ccl in + let Sigma (c, _, _) = redfun.e_redfun env (Proofview.Goal.sigma gl) (EConstr.of_constr ccl) in aux c | x -> x in @@ -1228,7 +1229,7 @@ let cut c = try (** Backward compat: ensure that [c] is well-typed. *) let typ = Typing.unsafe_type_of env sigma c in - let typ = whd_all env sigma typ in + let typ = whd_all env sigma (EConstr.of_constr typ) in match kind_of_term typ with | Sort _ -> true | _ -> false @@ -1237,7 +1238,7 @@ let cut c = if is_sort then let id = next_name_away_with_default "H" Anonymous (Tacmach.New.pf_ids_of_hyps gl) in (** Backward compat: normalize [c]. *) - let c = if normalize_cut then local_strong whd_betaiota sigma c else c in + let c = if normalize_cut then local_strong whd_betaiota sigma (EConstr.of_constr c) else c in Refine.refine ~unsafe:true { run = begin fun h -> let Sigma (f, h, p) = Evarutil.new_evar ~principal:true env h (mkArrow c (Vars.lift 1 concl)) in let Sigma (x, h, q) = Evarutil.new_evar env h c in @@ -1591,12 +1592,12 @@ let make_projection env sigma params cstr sign elim i n c u = noccur_between 1 (n-i-1) t (* to avoid surprising unifications, excludes flexible projection types or lambda which will be instantiated by Meta/Evar *) - && not (isEvar (fst (whd_betaiota_stack sigma t))) + && not (EConstr.isEvar sigma (fst (whd_betaiota_stack sigma (EConstr.of_constr t)))) && (accept_universal_lemma_under_conjunctions () || not (isRel t)) then let t = lift (i+1-n) t in - let abselim = beta_applist (elim,params@[t;branch]) in - let c = beta_applist (abselim, [mkApp (c, Context.Rel.to_extended_vect 0 sign)]) in + let abselim = beta_applist sigma (EConstr.of_constr elim, List.map EConstr.of_constr (params@[t;branch])) in + let c = beta_applist sigma (EConstr.of_constr abselim, [EConstr.of_constr (mkApp (c, Context.Rel.to_extended_vect 0 sign))]) in Some (it_mkLambda_or_LetIn c sign, it_mkProd_or_LetIn t sign) else None @@ -1630,7 +1631,7 @@ let descend_in_conjunctions avoid tac (err, info) c = | Some (_,_,isrec) -> let n = (constructors_nrealargs ind).(0) in let sort = Tacticals.New.elimination_sort_of_goal gl in - let IndType (indf,_) = find_rectype env sigma ccl in + let IndType (indf,_) = find_rectype env sigma (EConstr.of_constr ccl) in let (_,inst), params = dest_ind_family indf in let cstr = (get_constructors env indf).(0) in let elim = @@ -1703,7 +1704,7 @@ let general_apply with_delta with_destruct with_evars clear_flag (loc,(c,lbind)) let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in - let thm_ty0 = nf_betaiota sigma (Retyping.get_type_of env sigma c) in + let thm_ty0 = nf_betaiota sigma (EConstr.of_constr (Retyping.get_type_of env sigma c)) in let try_apply thm_ty nprod = try let n = nb_prod_modulo_zeta sigma (EConstr.of_constr thm_ty) - nprod in @@ -1716,7 +1717,7 @@ let general_apply with_delta with_destruct with_evars clear_flag (loc,(c,lbind)) let rec try_red_apply thm_ty (exn0, info) = try (* Try to head-reduce the conclusion of the theorem *) - let red_thm = try_red_product env sigma thm_ty in + let red_thm = try_red_product env sigma (EConstr.of_constr thm_ty) in tclORELSEOPT (try_apply red_thm concl_nprod) (function (e, info) -> match e with @@ -1829,7 +1830,7 @@ let progress_with_clause flags innerclause clause = with Not_found -> error "Unable to unify." let apply_in_once_main flags innerclause env sigma (d,lbind) = - let thm = nf_betaiota sigma (Retyping.get_type_of env sigma d) in + let thm = nf_betaiota sigma (EConstr.of_constr (Retyping.get_type_of env sigma d)) in let rec aux clause = try progress_with_clause flags innerclause clause with e when CErrors.noncritical e -> @@ -2127,7 +2128,7 @@ let apply_type newcl args = let env = Proofview.Goal.env gl in let store = Proofview.Goal.extra gl in Refine.refine { run = begin fun sigma -> - let newcl = nf_betaiota (Sigma.to_evar_map sigma) newcl (* As in former Logic.refine *) in + let newcl = nf_betaiota (Sigma.to_evar_map sigma) (EConstr.of_constr newcl) (* As in former Logic.refine *) in let Sigma (ev, sigma, p) = Evarutil.new_evar env sigma ~principal:true ~store newcl in Sigma (applist (ev, args), sigma, p) @@ -2318,7 +2319,7 @@ let rewrite_hyp_then assert_style with_evars thin l2r id tac = let sigma = Tacmach.New.project gl in let type_of = Tacmach.New.pf_unsafe_type_of gl in let whd_all = Tacmach.New.pf_apply whd_all gl in - let t = whd_all (type_of (mkVar id)) in + let t = whd_all (EConstr.of_constr (type_of (mkVar id))) in let eqtac, thin = match match_with_equality_type sigma t with | Some (hdcncl,[_;lhs;rhs]) -> if l2r && isVar lhs && not (occur_var env sigma (destVar lhs) (EConstr.of_constr rhs)) then @@ -2905,13 +2906,13 @@ let specialize (c,lbind) ipat = let clause = make_clenv_binding env sigma (c,Retyping.get_type_of env sigma c) lbind in let flags = { (default_unify_flags ()) with resolve_evars = true } in let clause = clenv_unify_meta_types ~flags clause in - let (thd,tstack) = whd_nored_stack clause.evd (clenv_value clause) in + let (thd,tstack) = whd_nored_stack clause.evd (EConstr.of_constr (clenv_value clause)) in let rec chk = function | [] -> [] - | t::l -> if occur_meta clause.evd (EConstr.of_constr t) then [] else t :: chk l + | t::l -> if occur_meta clause.evd t then [] else EConstr.Unsafe.to_constr t :: chk l in let tstack = chk tstack in - let term = applist(thd,List.map (nf_evar clause.evd) tstack) in + let term = applist(EConstr.Unsafe.to_constr thd,List.map (nf_evar clause.evd) tstack) in if occur_meta clause.evd (EConstr.of_constr term) then user_err (str "Cannot infer an instance for " ++ @@ -2964,7 +2965,7 @@ let unfold_body x = in Tacticals.New.afterHyp x begin fun aft -> let hl = List.fold_right (fun decl cl -> (NamedDecl.get_id decl, InHyp) :: cl) aft [] in - let rfun _ _ c = replace_vars [x, xval] c in + let rfun _ _ c = replace_vars [x, xval] (EConstr.Unsafe.to_constr c) in let reducth h = reduct_in_hyp rfun h in let reductc = reduct_in_concl (rfun, DEFAULTcast) in Tacticals.New.tclTHENLIST [Tacticals.New.tclMAP reducth hl; reductc] @@ -3519,7 +3520,7 @@ let decompose_indapp f args = let mk_term_eq env sigma ty t ty' t' = let sigma = Sigma.to_evar_map sigma in - if Reductionops.is_conv env sigma ty ty' then + if Reductionops.is_conv env sigma (EConstr.of_constr ty) (EConstr.of_constr ty') then mkEq ty t t', mkRefl ty' t' else mkHEq ty t ty' t', mkHRefl ty' t' @@ -3615,7 +3616,7 @@ let abstract_args gl generalize_vars dep id defined f args = *) let aux (prod, ctx, ctxenv, c, args, eqs, refls, nongenvars, vars, env) arg = let name, ty, arity = - let rel, c = Reductionops.splay_prod_n env !sigma 1 prod in + let rel, c = Reductionops.splay_prod_n env !sigma 1 (EConstr.of_constr prod) in let decl = List.hd rel in RelDecl.get_name decl, RelDecl.get_type decl, c in @@ -3765,8 +3766,8 @@ let specialize_eqs id gl = in let ty' = it_mkProd_or_LetIn ty ctx'' in let acc' = it_mkLambda_or_LetIn acc ctx'' in - let ty' = Tacred.whd_simpl env !evars ty' - and acc' = Tacred.whd_simpl env !evars acc' in + let ty' = Tacred.whd_simpl env !evars (EConstr.of_constr ty') + and acc' = Tacred.whd_simpl env !evars (EConstr.of_constr acc') in let ty' = Evarutil.nf_evar !evars ty' in if worked then tclTHENFIRST (Tacmach.internal_cut true id ty') @@ -4244,7 +4245,7 @@ let use_bindings env sigma elim must_be_closed (c,lbind) typ = known only by pattern-matching, as in the case of a term of the form "nat_rect ?A ?o ?s n", with ?A to be inferred by matching. *) - let sign,t = splay_prod env sigma typ in it_mkProd t sign + let sign,t = splay_prod env sigma (EConstr.of_constr typ) in it_mkProd t sign else (* Otherwise, we exclude the case of an induction argument in an explicitly functional type. Henceforth, we can complete the @@ -4261,14 +4262,14 @@ let use_bindings env sigma elim must_be_closed (c,lbind) typ = let (sigma, c) = pose_all_metas_as_evars env indclause.evd (clenv_value indclause) in Sigma.Unsafe.of_pair (c, sigma) with e when catchable_exception e -> - try find_clause (try_red_product env sigma typ) + try find_clause (try_red_product env sigma (EConstr.of_constr typ)) with Redelimination -> raise e in find_clause typ let check_expected_type env sigma (elimc,bl) elimt = (* Compute the expected template type of the term in case a using clause is given *) - let sign,_ = splay_prod env sigma elimt in + let sign,_ = splay_prod env sigma (EConstr.of_constr elimt) in let n = List.length sign in if n == 0 then error "Scheme cannot be applied."; let sigma,cl = make_evar_clause env sigma ~len:(n - 1) elimt in @@ -4283,7 +4284,7 @@ let check_enough_applied env sigma elim = | None -> (* No eliminator given *) fun u -> - let t,_ = decompose_app (whd_all env sigma u) in isInd t + let t,_ = decompose_app (whd_all env sigma (EConstr.of_constr u)) in isInd t | Some elimc -> let elimt = Retyping.get_type_of env sigma (fst elimc) in let scheme = compute_elim_sig ~elimc elimt in @@ -4604,7 +4605,7 @@ let maybe_betadeltaiota_concl allowred gl = if not allowred then concl else let env = Proofview.Goal.env gl in - whd_all env sigma concl + whd_all env sigma (EConstr.of_constr concl) let reflexivity_red allowred = Proofview.Goal.enter { enter = begin fun gl -> diff --git a/tactics/tactics.mli b/tactics/tactics.mli index 7acfb6286..268453152 100644 --- a/tactics/tactics.mli +++ b/tactics/tactics.mli @@ -128,7 +128,7 @@ val exact_proof : Constrexpr.constr_expr -> unit Proofview.tactic (** {6 Reduction tactics. } *) -type tactic_reduction = env -> evar_map -> constr -> constr +type tactic_reduction = env -> evar_map -> EConstr.t -> constr type change_arg = patvar_map -> constr Sigma.run diff --git a/toplevel/auto_ind_decl.ml b/toplevel/auto_ind_decl.ml index b1811d6a6..02c43aec5 100644 --- a/toplevel/auto_ind_decl.ml +++ b/toplevel/auto_ind_decl.ml @@ -180,7 +180,7 @@ let build_beq_scheme mode kn = let lifti = ndx in let rec aux c = let (c,a) = Reductionops.whd_betaiota_stack Evd.empty c in - match kind_of_term c with + match EConstr.kind Evd.empty (** FIXME *) c with | Rel x -> mkRel (x-nlist+ndx), Safe_typing.empty_private_constants | Var x -> let eid = id_of_string ("eq_"^(string_of_id x)) in @@ -189,7 +189,7 @@ let build_beq_scheme mode kn = with Not_found -> raise (ParameterWithoutEquality (VarRef x)) in mkVar eid, Safe_typing.empty_private_constants - | Cast (x,_,_) -> aux (applist (x,a)) + | Cast (x,_,_) -> aux (EConstr.applist (x,a)) | App _ -> assert false | Ind ((kn',i as ind'),u) (*FIXME: universes *) -> if eq_mind kn kn' then mkRel(eqA-nlist-i+nb_ind-1), Safe_typing.empty_private_constants @@ -205,7 +205,7 @@ let build_beq_scheme mode kn = in let args = Array.append - (Array.of_list (List.map (fun x -> lift lifti x) a)) eqa in + (Array.of_list (List.map (fun x -> lift lifti (EConstr.Unsafe.to_constr x)) a)) eqa in if Int.equal (Array.length args) 0 then eq, eff else mkApp (eq, args), eff with Not_found -> raise(EqNotFound (ind', fst ind)) @@ -217,7 +217,7 @@ let build_beq_scheme mode kn = | Const kn -> (match Environ.constant_opt_value_in env kn with | None -> raise (ParameterWithoutEquality (ConstRef (fst kn))) - | Some c -> aux (applist (c,a))) + | Some c -> aux (EConstr.applist (EConstr.of_constr c,a))) | Proj _ -> raise (EqUnknown "Proj") | Construct _ -> raise (EqUnknown "Construct") | Case _ -> raise (EqUnknown "Case") @@ -261,7 +261,7 @@ let build_beq_scheme mode kn = nparrec (nparrec+3+2*nb_cstr_args) (nb_cstr_args+ndx+1) - cc + (EConstr.of_constr cc) in eff := Safe_typing.concat_private eff' !eff; Array.set eqs ndx diff --git a/toplevel/class.ml b/toplevel/class.ml index 46854e5c0..5aa000b16 100644 --- a/toplevel/class.ml +++ b/toplevel/class.ml @@ -62,7 +62,7 @@ let explain_coercion_error g = function (* Verifications pour l'ajout d'une classe *) let check_reference_arity ref = - if not (Reductionops.is_arity (Global.env()) Evd.empty (Global.type_of_global_unsafe ref)) then + if not (Reductionops.is_arity (Global.env()) Evd.empty (EConstr.of_constr (Global.type_of_global_unsafe ref))) (** FIXME *) then raise (CoercionError (NotAClass ref)) let check_arity = function @@ -206,7 +206,7 @@ let build_id_coercion idf_opt source poly = let _ = if not (Reductionops.is_conv_leq env sigma - (Typing.unsafe_type_of env sigma val_f) typ_f) + (EConstr.of_constr (Typing.unsafe_type_of env sigma val_f)) (EConstr.of_constr typ_f)) then user_err (strbrk "Cannot be defined as coercion (maybe a bad number of arguments).") diff --git a/toplevel/command.ml b/toplevel/command.ml index 62bbd4b97..249d41845 100644 --- a/toplevel/command.ml +++ b/toplevel/command.ml @@ -45,7 +45,7 @@ let do_universe poly l = Declare.do_universe poly l let do_constraint poly l = Declare.do_constraint poly l let rec under_binders env sigma f n c = - if Int.equal n 0 then f env sigma c else + if Int.equal n 0 then f env sigma (EConstr.of_constr c) else match kind_of_term c with | Lambda (x,t,c) -> mkLambda (x,t,under_binders (push_rel (LocalAssum (x,t)) env) sigma f (n-1) c) @@ -395,7 +395,7 @@ let check_all_names_different indl = | _ -> raise (InductiveError (SameNamesOverlap l)) let mk_mltype_data evdref env assums arity indname = - let is_ml_type = is_sort env !evdref arity in + let is_ml_type = is_sort env !evdref (EConstr.of_constr arity) in (is_ml_type,indname,assums) let prepare_param = function @@ -961,10 +961,10 @@ let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation = (Printer.pr_constr_env env !evdref rel ++ str " is not an homogeneous binary relation.") in try - let ctx, ar = Reductionops.splay_prod_n env !evdref 2 relty in + let ctx, ar = Reductionops.splay_prod_n env !evdref 2 (EConstr.of_constr relty) in match ctx, kind_of_term ar with | [LocalAssum (_,t); LocalAssum (_,u)], Sort (Prop Null) - when Reductionops.is_conv env !evdref t u -> t + when Reductionops.is_conv env !evdref (EConstr.of_constr t) (EConstr.of_constr u) -> t | _, _ -> error () with e when CErrors.noncritical e -> error () in diff --git a/toplevel/himsg.ml b/toplevel/himsg.ml index bfe86053e..0279ff0c5 100644 --- a/toplevel/himsg.ml +++ b/toplevel/himsg.ml @@ -79,7 +79,7 @@ let rec contract3' env a b c = function let j_nf_betaiotaevar sigma j = { uj_val = Evarutil.nf_evar sigma j.uj_val; - uj_type = Reductionops.nf_betaiota sigma j.uj_type } + uj_type = Reductionops.nf_betaiota sigma (EConstr.of_constr j.uj_type) } let jv_nf_betaiotaevar sigma jl = Array.map (j_nf_betaiotaevar sigma) jl @@ -335,7 +335,7 @@ let explain_unification_error env sigma p1 p2 = function let explain_actual_type env sigma j t reason = let env = make_all_name_different env in let j = j_nf_betaiotaevar sigma j in - let t = Reductionops.nf_betaiota sigma t in + let t = Reductionops.nf_betaiota sigma (EConstr.of_constr t) in (** Actually print *) let pe = pr_ne_context_of (str "In environment") env sigma in let pc = pr_lconstr_env env sigma (Environ.j_val j) in @@ -351,7 +351,7 @@ let explain_actual_type env sigma j t reason = let explain_cant_apply_bad_type env sigma (n,exptyp,actualtyp) rator randl = let randl = jv_nf_betaiotaevar sigma randl in let exptyp = Evarutil.nf_evar sigma exptyp in - let actualtyp = Reductionops.nf_betaiota sigma actualtyp in + let actualtyp = Reductionops.nf_betaiota sigma (EConstr.of_constr actualtyp) in let rator = Evarutil.j_nf_evar sigma rator in let env = make_all_name_different env in let actualtyp, exptyp = pr_explicit env sigma actualtyp exptyp in diff --git a/toplevel/obligations.ml b/toplevel/obligations.ml index 1d5e4a2fa..8aee9d241 100644 --- a/toplevel/obligations.ml +++ b/toplevel/obligations.ml @@ -264,7 +264,7 @@ let pperror cmd = CErrors.user_err ~hdr:"Program" cmd let error s = pperror (str s) let reduce c = - Reductionops.clos_norm_flags CClosure.betaiota (Global.env ()) Evd.empty c + Reductionops.clos_norm_flags CClosure.betaiota (Global.env ()) Evd.empty (EConstr.of_constr c) exception NoObligations of Id.t option @@ -534,8 +534,8 @@ let declare_mutual_definition l = List.split3 (List.map (fun x -> let subs, typ = (subst_body true x) in - let term = snd (Reductionops.splay_lam_n (Global.env ()) Evd.empty len subs) in - let typ = snd (Reductionops.splay_prod_n (Global.env ()) Evd.empty len typ) in + let term = snd (Reductionops.splay_lam_n (Global.env ()) Evd.empty len (EConstr.of_constr subs)) in + let typ = snd (Reductionops.splay_prod_n (Global.env ()) Evd.empty len (EConstr.of_constr typ)) in x.prg_reduce term, x.prg_reduce typ, x.prg_implicits) l) in (* let fixdefs = List.map reduce_fix fixdefs in *) diff --git a/toplevel/record.ml b/toplevel/record.ml index 7dee4aae0..ffe7980ef 100644 --- a/toplevel/record.ml +++ b/toplevel/record.ml @@ -120,7 +120,7 @@ let typecheck_params_and_fields def id pl t ps nots fs = match t with | CSort (_, Misctypes.GType []) -> true | _ -> false in let s = interp_type_evars env evars ~impls:empty_internalization_env t in - let sred = Reductionops.whd_all env !evars s in + let sred = Reductionops.whd_all env !evars (EConstr.of_constr s) in (match kind_of_term sred with | Sort s' -> (if poly then diff --git a/toplevel/vernacentries.ml b/toplevel/vernacentries.ml index bbbdbdb67..593aa9578 100644 --- a/toplevel/vernacentries.ml +++ b/toplevel/vernacentries.ml @@ -1586,7 +1586,7 @@ let vernac_check_may_eval redexp glopt rc = match redexp with | None -> let l = Evar.Set.union (Evd.evars_of_term j.Environ.uj_val) (Evd.evars_of_term j.Environ.uj_type) in - let j = { j with Environ.uj_type = Reductionops.nf_betaiota sigma' j.Environ.uj_type } in + let j = { j with Environ.uj_type = Reductionops.nf_betaiota sigma' (EConstr.of_constr j.Environ.uj_type) } in Feedback.msg_notice (print_judgment env sigma' j ++ pr_ne_evar_set (fnl () ++ str "where" ++ fnl ()) (mt ()) sigma' l ++ Printer.pr_universe_ctx sigma uctx) -- cgit v1.2.3 From 2db085e62f7797cc999518eb58983ac059763e1f Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 4 Nov 2016 14:13:08 +0100 Subject: Vnorm API using EConstr. --- pretyping/vnorm.ml | 5 +++++ pretyping/vnorm.mli | 4 ++-- proofs/redexpr.ml | 7 ++----- 3 files changed, 9 insertions(+), 7 deletions(-) diff --git a/pretyping/vnorm.ml b/pretyping/vnorm.ml index 8c3de7cfd..60f99fd3d 100644 --- a/pretyping/vnorm.ml +++ b/pretyping/vnorm.ml @@ -352,6 +352,11 @@ and nf_cofix env sigma cf = mkCoFix (init,(name,cft,cfb)) let cbv_vm env sigma c t = + if Termops.occur_meta_or_existential sigma c then + CErrors.error "vm_compute does not support existential variables."; + (** This evar-normalizes terms beforehand *) + let c = EConstr.to_constr sigma c in + let t = EConstr.to_constr sigma t in let v = Vconv.val_of_constr env c in nf_val env sigma v t diff --git a/pretyping/vnorm.mli b/pretyping/vnorm.mli index bc6eec851..650f3f291 100644 --- a/pretyping/vnorm.mli +++ b/pretyping/vnorm.mli @@ -6,8 +6,8 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -open Term +open EConstr open Environ (** {6 Reduction functions } *) -val cbv_vm : env -> Evd.evar_map -> constr -> types -> constr +val cbv_vm : env -> Evd.evar_map -> constr -> types -> Constr.t diff --git a/proofs/redexpr.ml b/proofs/redexpr.ml index 40a8077a7..348cd1bcb 100644 --- a/proofs/redexpr.ml +++ b/proofs/redexpr.ml @@ -24,11 +24,8 @@ open Misctypes (* call by value normalisation function using the virtual machine *) let cbv_vm env sigma c = - let c = EConstr.Unsafe.to_constr c in - let ctyp = Retyping.get_type_of env sigma c in - if Termops.occur_meta_or_existential sigma (EConstr.of_constr c) then - error "vm_compute does not support existential variables."; - Vnorm.cbv_vm env sigma c ctyp + let ctyp = Retyping.get_type_of env sigma (EConstr.Unsafe.to_constr c) in + Vnorm.cbv_vm env sigma c (EConstr.of_constr ctyp) let warn_native_compute_disabled = CWarnings.create ~name:"native-compute-disabled" ~category:"native-compiler" -- cgit v1.2.3 From 6bd193ff409b01948751525ce0f905916d7a64bd Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 4 Nov 2016 14:38:35 +0100 Subject: Nativenorm API using EConstr. --- pretyping/nativenorm.ml | 6 +++++- pretyping/nativenorm.mli | 5 +++-- pretyping/pretyping.ml | 2 +- proofs/redexpr.ml | 5 ++--- 4 files changed, 11 insertions(+), 7 deletions(-) diff --git a/pretyping/nativenorm.ml b/pretyping/nativenorm.ml index e45c920a3..c8bcae0c8 100644 --- a/pretyping/nativenorm.ml +++ b/pretyping/nativenorm.ml @@ -377,7 +377,9 @@ let evars_of_evar_map sigma = Nativelambda.evars_typ = Evd.existential_type sigma; Nativelambda.evars_metas = Evd.meta_type sigma } -let native_norm env sigma c ty = +let native_norm env sigma c ty = + let c = EConstr.Unsafe.to_constr c in + let ty = EConstr.Unsafe.to_constr ty in if Coq_config.no_native_compiler then error "Native_compute reduction has been disabled at configure time." else @@ -407,5 +409,7 @@ let native_conv_generic pb sigma t = Nativeconv.native_conv_gen pb (evars_of_evar_map sigma) t let native_infer_conv ?(pb=Reduction.CUMUL) env sigma t1 t2 = + let t1 = EConstr.Unsafe.to_constr t1 in + let t2 = EConstr.Unsafe.to_constr t2 in Reductionops.infer_conv_gen (fun pb ~l2r sigma ts -> native_conv_generic pb sigma) ~catch_incon:true ~pb env sigma t1 t2 diff --git a/pretyping/nativenorm.mli b/pretyping/nativenorm.mli index 0b1ce8e51..ba46138a4 100644 --- a/pretyping/nativenorm.mli +++ b/pretyping/nativenorm.mli @@ -5,13 +5,14 @@ (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -open Term + +open EConstr open Environ open Evd (** This module implements normalization by evaluation to OCaml code *) -val native_norm : env -> evar_map -> constr -> types -> constr +val native_norm : env -> evar_map -> constr -> types -> Constr.t (** Conversion with inference of universe constraints *) val native_infer_conv : ?pb:conv_pb -> env -> evar_map -> constr -> constr -> diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml index 3a6d4f36c..2f42ad395 100644 --- a/pretyping/pretyping.ml +++ b/pretyping/pretyping.ml @@ -1021,7 +1021,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre let cj = pretype empty_tycon env evdref lvar c in let cty = nf_evar !evdref cj.uj_type and tval = nf_evar !evdref tval in begin - let (evd,b) = Nativenorm.native_infer_conv env.ExtraEnv.env !evdref cty tval in + let (evd,b) = Nativenorm.native_infer_conv env.ExtraEnv.env !evdref (EConstr.of_constr cty) (EConstr.of_constr tval) in if b then (evdref := evd; cj, tval) else error_actual_type ~loc env.ExtraEnv.env !evdref cj tval diff --git a/proofs/redexpr.ml b/proofs/redexpr.ml index 348cd1bcb..62fe2c17c 100644 --- a/proofs/redexpr.ml +++ b/proofs/redexpr.ml @@ -37,9 +37,8 @@ let cbv_native env sigma c = (warn_native_compute_disabled (); cbv_vm env sigma c) else - let c = EConstr.Unsafe.to_constr c in - let ctyp = Retyping.get_type_of env sigma c in - Nativenorm.native_norm env sigma c ctyp + let ctyp = Retyping.get_type_of env sigma (EConstr.Unsafe.to_constr c) in + Nativenorm.native_norm env sigma c (EConstr.of_constr ctyp) let whd_cbn flags env sigma t = let (state,_) = -- cgit v1.2.3 From d528fdaf12b74419c47698cca7c6f1ec762245a3 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 4 Nov 2016 14:48:36 +0100 Subject: Retyping API using EConstr. --- ltac/extratactics.ml4 | 4 +- ltac/rewrite.ml | 38 +++++------ plugins/extraction/extraction.ml | 6 +- plugins/micromega/coq_micromega.ml | 2 +- plugins/rtauto/refl_tauto.ml | 8 +-- plugins/setoid_ring/newring.ml | 14 ++--- pretyping/cases.ml | 14 ++--- pretyping/classops.ml | 2 +- pretyping/coercion.ml | 4 +- pretyping/constr_matching.ml | 8 +-- pretyping/detyping.ml | 6 +- pretyping/evarconv.ml | 12 ++-- pretyping/evarsolve.ml | 16 ++--- pretyping/evarsolve.mli | 2 +- pretyping/inductiveops.ml | 4 +- pretyping/inductiveops.mli | 2 +- pretyping/patternops.ml | 2 +- pretyping/pretyping.ml | 8 +-- pretyping/retyping.ml | 126 ++++++++++++++++++++----------------- pretyping/retyping.mli | 16 ++--- pretyping/tacred.ml | 4 +- pretyping/typeclasses.ml | 4 +- pretyping/typing.ml | 2 +- pretyping/unification.ml | 32 +++++----- proofs/clenv.ml | 6 +- proofs/logic.ml | 17 ++--- proofs/redexpr.ml | 4 +- proofs/refine.ml | 2 +- proofs/tacmach.ml | 4 +- tactics/class_tactics.ml | 4 +- tactics/contradiction.ml | 2 +- tactics/eqschemes.ml | 2 +- tactics/equality.ml | 24 +++---- tactics/hints.ml | 2 +- tactics/inv.ml | 4 +- tactics/tacticals.ml | 8 +-- tactics/tactics.ml | 44 ++++++------- toplevel/command.ml | 4 +- toplevel/indschemes.ml | 2 +- toplevel/record.ml | 2 +- toplevel/vernacentries.ml | 2 +- 41 files changed, 240 insertions(+), 229 deletions(-) diff --git a/ltac/extratactics.ml4 b/ltac/extratactics.ml4 index 2405e3c42..ca1f6e638 100644 --- a/ltac/extratactics.ml4 +++ b/ltac/extratactics.ml4 @@ -298,7 +298,7 @@ let project_hint pri l2r r = let env = Global.env() in let sigma = Evd.from_env env in let sigma, c = Evd.fresh_global env sigma gr in - let t = Retyping.get_type_of env sigma c in + let t = Retyping.get_type_of env sigma (EConstr.of_constr c) in let t = Tacred.reduce_to_quantified_ref env sigma (Lazy.force coq_iff_ref) t in let sign,ccl = decompose_prod_assum t in @@ -655,7 +655,7 @@ let hResolve id c occ t = in let t_constr,ctx = resolve_hole (subst_var_with_hole occ id t_raw) in let sigma = Evd.merge_universe_context sigma ctx in - let t_constr_type = Retyping.get_type_of env sigma t_constr in + let t_constr_type = Retyping.get_type_of env sigma (EConstr.of_constr t_constr) in let tac = (change_concl (mkLetIn (Anonymous,t_constr,t_constr_type,concl))) in diff --git a/ltac/rewrite.ml b/ltac/rewrite.ml index 7ef3ace53..80307ce8e 100644 --- a/ltac/rewrite.ml +++ b/ltac/rewrite.ml @@ -431,7 +431,7 @@ module TypeGlobal = struct end let sort_of_rel env evm rel = - Reductionops.sort_of_arity env evm (EConstr.of_constr (Retyping.get_type_of env evm rel)) + Reductionops.sort_of_arity env evm (EConstr.of_constr (Retyping.get_type_of env evm (EConstr.of_constr rel))) let is_applied_rewrite_relation = PropGlobal.is_applied_rewrite_relation @@ -499,20 +499,20 @@ let rec decompose_app_rel env evd t = let decompose_app_rel env evd t = let (rel, t1, t2) = decompose_app_rel env evd t in - let ty = Retyping.get_type_of env evd rel in + let ty = Retyping.get_type_of env evd (EConstr.of_constr rel) in let () = if not (Reduction.is_arity env ty) then error_no_relation () in (rel, t1, t2) let decompose_applied_relation env sigma (c,l) = let open Context.Rel.Declaration in - let ctype = Retyping.get_type_of env sigma c in + let ctype = Retyping.get_type_of env sigma (EConstr.of_constr c) in let find_rel ty = let sigma, cl = Clenv.make_evar_clause env sigma ty in let sigma = Clenv.solve_evar_clause env sigma true cl l in let { Clenv.cl_holes = holes; Clenv.cl_concl = t } = cl in let (equiv, c1, c2) = decompose_app_rel env sigma t in - let ty1 = Retyping.get_type_of env sigma c1 in - let ty2 = Retyping.get_type_of env sigma c2 in + let ty1 = Retyping.get_type_of env sigma (EConstr.of_constr c1) in + let ty2 = Retyping.get_type_of env sigma (EConstr.of_constr c2) in match evd_convertible env sigma ty1 ty2 with | None -> None | Some sigma -> @@ -719,8 +719,8 @@ let unify_eqn (car, rel, prf, c1, c2, holes, sort) l2r flags env (sigma, cstrs) let c1 = nf c1 and c2 = nf c2 and rew_car = nf car and rel = nf rel and prf = nf prf in - let ty1 = Retyping.get_type_of env evd c1 in - let ty2 = Retyping.get_type_of env evd c2 in + let ty1 = Retyping.get_type_of env evd (EConstr.of_constr c1) in + let ty2 = Retyping.get_type_of env evd (EConstr.of_constr c2) in let () = if not (convertible env evd ty2 ty1) then raise Reduction.NotConvertible in let rew_evars = evd, cstrs in let rew_prf = RewPrf (rel, prf) in @@ -923,15 +923,15 @@ let reset_env env = let fold_match ?(force=false) env sigma c = let (ci, p, c, brs) = destCase c in - let cty = Retyping.get_type_of env sigma c in + let cty = Retyping.get_type_of env sigma (EConstr.of_constr c) in let dep, pred, exists, (sk,eff) = let env', ctx, body = let ctx, pred = decompose_lam_assum p in let env' = Environ.push_rel_context ctx env in env', ctx, pred in - let sortp = Retyping.get_sort_family_of env' sigma body in - let sortc = Retyping.get_sort_family_of env sigma cty in + let sortp = Retyping.get_sort_family_of env' sigma (EConstr.of_constr body) in + let sortc = Retyping.get_sort_family_of env sigma (EConstr.of_constr cty) in let dep = not (noccurn 1 body) in let pred = if dep then p else it_mkProd_or_LetIn (subst1 mkProp body) (List.tl ctx) @@ -985,7 +985,7 @@ let subterm all flags (s : 'a pure_strategy) : 'a pure_strategy = if not (Option.is_empty progress) && not all then state, (None :: acc, evars, progress) else - let argty = Retyping.get_type_of env (goalevars evars) arg in + let argty = Retyping.get_type_of env (goalevars evars) (EConstr.of_constr arg) in let state, res = s.strategy { state ; env ; unfresh ; term1 = arg ; ty1 = argty ; @@ -1033,7 +1033,7 @@ let subterm all flags (s : 'a pure_strategy) : 'a pure_strategy = in state, res in if flags.on_morphisms then - let mty = Retyping.get_type_of env (goalevars evars) m in + let mty = Retyping.get_type_of env (goalevars evars) (EConstr.of_constr m) in let evars, cstr', m, mty, argsl, args = let argsl = Array.to_list args in let lift = if prop then PropGlobal.lift_cstr else TypeGlobal.lift_cstr in @@ -1075,8 +1075,8 @@ let subterm all flags (s : 'a pure_strategy) : 'a pure_strategy = | Prod (n, x, b) when noccurn 1 b -> let b = subst1 mkProp b in - let tx = Retyping.get_type_of env (goalevars evars) x - and tb = Retyping.get_type_of env (goalevars evars) b in + let tx = Retyping.get_type_of env (goalevars evars) (EConstr.of_constr x) + and tb = Retyping.get_type_of env (goalevars evars) (EConstr.of_constr b) in let arr = if prop then PropGlobal.arrow_morphism else TypeGlobal.arrow_morphism in @@ -1154,7 +1154,7 @@ let subterm all flags (s : 'a pure_strategy) : 'a pure_strategy = let n' = name_app (fun id -> Tactics.fresh_id_in_env unfresh id env) n in let open Context.Rel.Declaration in let env' = Environ.push_rel (LocalAssum (n', t)) env in - let bty = Retyping.get_type_of env' (goalevars evars) b in + let bty = Retyping.get_type_of env' (goalevars evars) (EConstr.of_constr b) in let unlift = if prop then PropGlobal.unlift_cstr else TypeGlobal.unlift_cstr in let state, b' = s.strategy { state ; env = env' ; unfresh ; term1 = b ; ty1 = bty ; @@ -1181,7 +1181,7 @@ let subterm all flags (s : 'a pure_strategy) : 'a pure_strategy = in state, res | Case (ci, p, c, brs) -> - let cty = Retyping.get_type_of env (goalevars evars) c in + let cty = Retyping.get_type_of env (goalevars evars) (EConstr.of_constr c) in let evars', eqty = app_poly_sort prop env evars coq_eq [| cty |] in let cstr' = Some eqty in let state, c' = s.strategy { state ; env ; unfresh ; @@ -1457,7 +1457,7 @@ let rewrite_with l2r flags c occs : strategy = { strategy = } let apply_strategy (s : strategy) env unfresh concl (prop, cstr) evars = - let ty = Retyping.get_type_of env (goalevars evars) concl in + let ty = Retyping.get_type_of env (goalevars evars) (EConstr.of_constr concl) in let _, res = s.strategy { state = () ; env ; unfresh ; term1 = concl ; ty1 = ty ; cstr = (prop, Some cstr) ; evars } in @@ -1868,7 +1868,7 @@ let declare_projection n instance_id r = let env = Global.env () in let sigma = Evd.from_env env in let sigma,c = Evd.fresh_global env sigma r in - let ty = Retyping.get_type_of env sigma c in + let ty = Retyping.get_type_of env sigma (EConstr.of_constr c) in let term = proper_projection c ty in let sigma, typ = Typing.type_of env sigma term in let ctx, typ = decompose_prod_assum typ in @@ -2060,7 +2060,7 @@ let unification_rewrite l2r c1 c2 sigma prf car rel but env = and car = nf car and rel = nf rel in check_evar_map_of_evars_defs sigma; let prf = nf prf in - let prfty = nf (Retyping.get_type_of env sigma prf) in + let prfty = nf (Retyping.get_type_of env sigma (EConstr.of_constr prf)) in let sort = sort_of_rel env sigma but in let abs = prf, prfty in let prf = mkRel 1 in diff --git a/plugins/extraction/extraction.ml b/plugins/extraction/extraction.ml index 6ca34036a..42a8cac69 100644 --- a/plugins/extraction/extraction.ml +++ b/plugins/extraction/extraction.ml @@ -42,11 +42,11 @@ let none = Evd.empty let type_of env c = let polyprop = (lang() == Haskell) in - Retyping.get_type_of ~polyprop env none (strip_outer_cast none (EConstr.of_constr c)) + Retyping.get_type_of ~polyprop env none (EConstr.of_constr (strip_outer_cast none (EConstr.of_constr c))) let sort_of env c = let polyprop = (lang() == Haskell) in - Retyping.get_sort_family_of ~polyprop env none (strip_outer_cast none (EConstr.of_constr c)) + Retyping.get_sort_family_of ~polyprop env none (EConstr.of_constr (strip_outer_cast none (EConstr.of_constr c))) (*S Generation of flags and signatures. *) @@ -595,7 +595,7 @@ let rec extract_term env mle mlt c args = | Construct (cp,_) -> extract_cons_app env mle mlt cp args | Proj (p, c) -> - let term = Retyping.expand_projection env (Evd.from_env env) p c [] in + let term = Retyping.expand_projection env (Evd.from_env env) p (EConstr.of_constr c) [] in extract_term env mle mlt term args | Rel n -> (* As soon as the expected [mlt] for the head is known, *) diff --git a/plugins/micromega/coq_micromega.ml b/plugins/micromega/coq_micromega.ml index 9fb1463db..a943ef2b0 100644 --- a/plugins/micromega/coq_micromega.ml +++ b/plugins/micromega/coq_micromega.ml @@ -1199,7 +1199,7 @@ struct with e when CErrors.noncritical e -> (X(t),env,tg) in let is_prop term = - let sort = Retyping.get_sort_of (Tacmach.pf_env gl) (Tacmach.project gl) term in + let sort = Retyping.get_sort_of (Tacmach.pf_env gl) (Tacmach.project gl) (EConstr.of_constr term) in Term.is_prop_sort sort in let rec xparse_formula env tg term = diff --git a/plugins/rtauto/refl_tauto.ml b/plugins/rtauto/refl_tauto.ml index b129b0bb3..f88b3a700 100644 --- a/plugins/rtauto/refl_tauto.ml +++ b/plugins/rtauto/refl_tauto.ml @@ -94,9 +94,9 @@ let rec make_form atom_env gls term = let cciterm=special_whd gls term in match kind_of_term cciterm with Prod(_,a,b) -> - if EConstr.Vars.noccurn Evd.empty (** FIXME *) 1 (EConstr.of_constr b) && + if EConstr.Vars.noccurn (Tacmach.project gls) 1 (EConstr.of_constr b) && Retyping.get_sort_family_of - (pf_env gls) (Tacmach.project gls) a == InProp + (pf_env gls) (Tacmach.project gls) (EConstr.of_constr a) == InProp then let fa=make_form atom_env gls a in let fb=make_form atom_env gls b in @@ -136,7 +136,7 @@ let rec make_hyps atom_env gls lenv = function make_hyps atom_env gls (typ::lenv) rest in if List.exists (fun c -> Termops.local_occur_var Evd.empty (** FIXME *) id (EConstr.of_constr c)) lenv || (Retyping.get_sort_family_of - (pf_env gls) (Tacmach.project gls) typ != InProp) + (pf_env gls) (Tacmach.project gls) (EConstr.of_constr typ) != InProp) then hrec else @@ -262,7 +262,7 @@ let rtauto_tac gls= let gl=pf_concl gls in let _= if Retyping.get_sort_family_of - (pf_env gls) (Tacmach.project gls) gl != InProp + (pf_env gls) (Tacmach.project gls) (EConstr.of_constr gl) != InProp then user_err ~hdr:"rtauto" (Pp.str "goal should be in Prop") in let glf=make_form gamma gls gl in let hyps=make_hyps gamma gls [gl] (pf_hyps gls) in diff --git a/plugins/setoid_ring/newring.ml b/plugins/setoid_ring/newring.ml index cf0f51911..e1b95ddbc 100644 --- a/plugins/setoid_ring/newring.ml +++ b/plugins/setoid_ring/newring.ml @@ -345,9 +345,9 @@ let ring_for_carrier r = Cmap.find r !from_carrier let find_ring_structure env sigma l = match l with | t::cl' -> - let ty = Retyping.get_type_of env sigma t in + let ty = Retyping.get_type_of env sigma (EConstr.of_constr t) in let check c = - let ty' = Retyping.get_type_of env sigma c in + let ty' = Retyping.get_type_of env sigma (EConstr.of_constr c) in if not (Reductionops.is_conv env sigma (EConstr.of_constr ty) (EConstr.of_constr ty')) then user_err ~hdr:"ring" (str"arguments of ring_simplify do not have all the same type") @@ -540,7 +540,7 @@ let build_setoid_params env evd r add mul opp req eqth = | None -> ring_equality env evd (r,add,mul,opp,req) let dest_ring env sigma th_spec = - let th_typ = Retyping.get_type_of env sigma th_spec in + let th_typ = Retyping.get_type_of env sigma (EConstr.of_constr th_spec) in match kind_of_term th_typ with App(f,[|r;zero;one;add;mul;sub;opp;req|]) when eq_constr_nounivs f (Lazy.force coq_almost_ring_theory) -> @@ -571,7 +571,7 @@ let interp_cst_tac env sigma rk kind (zero,one,add,mul,opp) cst_tac = TacArg(Loc.ghost,TacCall(Loc.ghost,t,[])) let make_hyp env evd c = - let t = Retyping.get_type_of env !evd c in + let t = Retyping.get_type_of env !evd (EConstr.of_constr c) in plapp evd coq_mkhypo [|t;c|] let make_hyp_list env evd lH = @@ -796,7 +796,7 @@ let af_ar = my_reference"AF_AR" let f_r = my_reference"F_R" let sf_sr = my_reference"SF_SR" let dest_field env evd th_spec = - let th_typ = Retyping.get_type_of env !evd th_spec in + let th_typ = Retyping.get_type_of env !evd (EConstr.of_constr th_spec) in match kind_of_term th_typ with | App(f,[|r;zero;one;add;mul;sub;opp;div;inv;req|]) when is_global (Lazy.force afield_theory) f -> @@ -825,9 +825,9 @@ let find_field_structure env sigma l = check_required_library (cdir@["Field_tac"]); match l with | t::cl' -> - let ty = Retyping.get_type_of env sigma t in + let ty = Retyping.get_type_of env sigma (EConstr.of_constr t) in let check c = - let ty' = Retyping.get_type_of env sigma c in + let ty' = Retyping.get_type_of env sigma (EConstr.of_constr c) in if not (Reductionops.is_conv env sigma (EConstr.of_constr ty) (EConstr.of_constr ty')) then user_err ~hdr:"field" (str"arguments of field_simplify do not have all the same type") diff --git a/pretyping/cases.ml b/pretyping/cases.ml index be72091a9..4dd502106 100644 --- a/pretyping/cases.ml +++ b/pretyping/cases.ml @@ -1480,7 +1480,7 @@ and compile_alias initial pb (na,orig,(expanded,expanded_typ)) rest = if not (Flags.is_program_mode ()) && (isRel orig || isVar orig) then (* Try to compile first using non expanded alias *) try - if initial then f orig (Retyping.get_type_of pb.env !(pb.evdref) orig) + if initial then f orig (Retyping.get_type_of pb.env !(pb.evdref) (EConstr.of_constr orig)) else just_pop () with e when precatchable_exception e -> (* Try then to compile using expanded alias *) @@ -1495,7 +1495,7 @@ and compile_alias initial pb (na,orig,(expanded,expanded_typ)) rest = (* Could be needed in case of a recursive call which requires to be on a variable for size reasons *) pb.evdref := sigma; - if initial then f orig (Retyping.get_type_of pb.env !(pb.evdref) orig) + if initial then f orig (Retyping.get_type_of pb.env !(pb.evdref) (EConstr.of_constr orig)) else just_pop () @@ -1627,7 +1627,7 @@ let abstract_tycon loc env evdref subst tycon extenv t = let t = whd_evar !evdref t in match kind_of_term t with | Rel n when is_local_def (lookup_rel n env) -> t | Evar ev -> - let ty = get_type_of env !evdref t in + let ty = get_type_of env !evdref (EConstr.of_constr t) in let ty = Evarutil.evd_comb1 (refresh_universes (Some false) env) evdref ty in let inst = List.map_i @@ -1649,7 +1649,7 @@ let abstract_tycon loc env evdref subst tycon extenv t = | (_, _, u) :: _ -> (* u is in extenv *) let vl = List.map pi1 good in let ty = - let ty = get_type_of env !evdref t in + let ty = get_type_of env !evdref (EConstr.of_constr t) in Evarutil.evd_comb1 (refresh_universes (Some false) env) evdref ty in let ty = lift (-k) (aux x ty) in @@ -1798,7 +1798,7 @@ let build_inversion_problem loc env sigma tms t = it = None } } ] in (* [pb] is the auxiliary pattern-matching serving as skeleton for the return type of the original problem Xi *) - let s' = Retyping.get_sort_of env sigma t in + let s' = Retyping.get_sort_of env sigma (EConstr.of_constr t) in let sigma, s = Evd.new_sort_variable univ_flexible_alg sigma in let sigma = Evd.set_leq_sort env sigma s' s in let evdref = ref sigma in @@ -2067,7 +2067,7 @@ let constr_of_pat env evdref arsign pat avoid = let cstr = mkConstructU ci.cs_cstr in let app = applistc cstr (List.map (lift (List.length sign)) params) in let app = applistc app args in - let apptype = Retyping.get_type_of env ( !evdref) app in + let apptype = Retyping.get_type_of env ( !evdref) (EConstr.of_constr app) in let IndType (indf, realargs) = find_rectype env (!evdref) (EConstr.of_constr apptype) in match alias with Anonymous -> @@ -2325,7 +2325,7 @@ let build_dependent_signature env evdref avoid tomatchs arsign = (fun (env, nargeqs, argeqs, refl_args, slift, argsign') arg decl -> let name = RelDecl.get_name decl in let t = RelDecl.get_type decl in - let argt = Retyping.get_type_of env !evdref arg in + let argt = Retyping.get_type_of env !evdref (EConstr.of_constr arg) in let eq, refl_arg = if Reductionops.is_conv env !evdref (EConstr.of_constr argt) (EConstr.of_constr t) then (mk_eq evdref (lift (nargeqs + slift) argt) diff --git a/pretyping/classops.ml b/pretyping/classops.ml index fd21f5bd1..577f41a7d 100644 --- a/pretyping/classops.ml +++ b/pretyping/classops.ml @@ -441,7 +441,7 @@ let cache_coercion (_, c) = let is, _ = class_info c.coercion_source in let it, _ = class_info c.coercion_target in let value, ctx = Universes.fresh_global_instance (Global.env()) c.coercion_type in - let typ = Retyping.get_type_of (Global.env ()) Evd.empty value in + let typ = Retyping.get_type_of (Global.env ()) Evd.empty (EConstr.of_constr value) in let xf = { coe_value = value; coe_type = typ; diff --git a/pretyping/coercion.ml b/pretyping/coercion.ml index b062da1f4..6a7f90463 100644 --- a/pretyping/coercion.ml +++ b/pretyping/coercion.ml @@ -65,7 +65,7 @@ let apply_coercion_args env evd check isproj argl funj = | h::restl -> (* On devrait pouvoir s'arranger pour qu'on n'ait pas a faire hnf_constr *) match kind_of_term (whd_all env evd (EConstr.of_constr typ)) with | Prod (_,c1,c2) -> - if check && not (e_cumul env evdref (Retyping.get_type_of env evd h) c1) then + if check && not (e_cumul env evdref (Retyping.get_type_of env evd (EConstr.of_constr h)) c1) then raise NoCoercion; apply_rec (h::acc) (subst1 h c2) restl | _ -> anomaly (Pp.str "apply_coercion_args") @@ -488,7 +488,7 @@ let rec inh_conv_coerce_to_fail loc env evd rigidonly v t c1 = let v2 = Option.map (fun v -> beta_applist evd' (EConstr.of_constr (lift 1 v),[EConstr.of_constr v1])) v in let t2 = match v2 with | None -> subst_term evd' (EConstr.of_constr v1) (EConstr.of_constr t2) - | Some v2 -> Retyping.get_type_of env1 evd' v2 in + | Some v2 -> Retyping.get_type_of env1 evd' (EConstr.of_constr v2) in let (evd'',v2') = inh_conv_coerce_to_fail loc env1 evd' rigidonly v2 t2 u2 in (evd'', Option.map (fun v2' -> mkLambda (name, u1, v2')) v2') | _ -> raise (NoCoercionNoUnifier (best_failed_evd,e)) diff --git a/pretyping/constr_matching.ml b/pretyping/constr_matching.ml index 66e690714..1261844a0 100644 --- a/pretyping/constr_matching.ml +++ b/pretyping/constr_matching.ml @@ -221,7 +221,7 @@ let matches_core env sigma convert allow_partial_app allow_bound_rels else (* Might be a projection on the right *) match kind_of_term c2 with | Proj (pr, c) when not (Projection.unfolded pr) -> - (try let term = Retyping.expand_projection env sigma pr c (Array.to_list args2) in + (try let term = Retyping.expand_projection env sigma pr (EConstr.of_constr c) (Array.map_to_list EConstr.of_constr args2) in sorec ctx env subst p term with Retyping.RetypeError _ -> raise PatternMatchingFailure) | _ -> raise PatternMatchingFailure) @@ -237,7 +237,7 @@ let matches_core env sigma convert allow_partial_app allow_bound_rels with Invalid_argument _ -> raise PatternMatchingFailure else raise PatternMatchingFailure | _, Proj (pr,c) when not (Projection.unfolded pr) -> - (try let term = Retyping.expand_projection env sigma pr c (Array.to_list arg2) in + (try let term = Retyping.expand_projection env sigma pr (EConstr.of_constr c) (Array.map_to_list EConstr.of_constr arg2) in sorec ctx env subst p term with Retyping.RetypeError _ -> raise PatternMatchingFailure) | _, _ -> @@ -249,7 +249,7 @@ let matches_core env sigma convert allow_partial_app allow_bound_rels raise PatternMatchingFailure | PApp (c, args), Proj (pr, c2) -> - (try let term = Retyping.expand_projection env sigma pr c2 [] in + (try let term = Retyping.expand_projection env sigma pr (EConstr.of_constr c2) [] in sorec ctx env subst p term with Retyping.RetypeError _ -> raise PatternMatchingFailure) @@ -440,7 +440,7 @@ let sub_match ?(partial_app=false) ?(closed=true) env sigma pat c = let next_mk_ctx le = mk_ctx (mkProj (p,List.hd le)) in if partial_app then try - let term = Retyping.expand_projection env sigma p c' [] in + let term = Retyping.expand_projection env sigma p (EConstr.of_constr c') [] in aux env term mk_ctx next with Retyping.RetypeError _ -> next () else diff --git a/pretyping/detyping.ml b/pretyping/detyping.ml index a4d943cfa..e5e778f23 100644 --- a/pretyping/detyping.ml +++ b/pretyping/detyping.ml @@ -501,7 +501,7 @@ let rec detype flags avoid env sigma t = try let pb = Environ.lookup_projection p (snd env) in let body = pb.Declarations.proj_body in - let ty = Retyping.get_type_of (snd env) sigma c in + let ty = Retyping.get_type_of (snd env) sigma (EConstr.of_constr c) in let ((ind,u), args) = Inductiveops.find_mrectype (snd env) sigma (EConstr.of_constr ty) in let body' = strip_lam_assum body in let body' = subst_instance_constr u body' in @@ -512,7 +512,7 @@ let rec detype flags avoid env sigma t = else if print_primproj_params () then try - let c = Retyping.expand_projection (snd env) sigma p c [] in + let c = Retyping.expand_projection (snd env) sigma p (EConstr.of_constr c) [] in detype flags avoid env sigma c with Retyping.RetypeError _ -> noparams () else noparams () @@ -689,7 +689,7 @@ and detype_binder (lax,isgoal as flags) bk avoid env sigma na body ty c = | BLetIn -> let c = detype (lax,false) avoid env sigma (Option.get body) in (* Heuristic: we display the type if in Prop *) - let s = try Retyping.get_sort_family_of (snd env) sigma ty with _ when !Flags.in_debugger || !Flags.in_toplevel -> InType (* Can fail because of sigma missing in debugger *) in + let s = try Retyping.get_sort_family_of (snd env) sigma (EConstr.of_constr ty) with _ when !Flags.in_debugger || !Flags.in_toplevel -> InType (* Can fail because of sigma missing in debugger *) in let c = if s != InProp then c else GCast (dl, c, CastConv (detype (lax,false) avoid env sigma ty)) in GLetIn (dl, na', c, r) diff --git a/pretyping/evarconv.ml b/pretyping/evarconv.ml index f54a57d57..47db71cc6 100644 --- a/pretyping/evarconv.ml +++ b/pretyping/evarconv.ml @@ -165,7 +165,7 @@ let check_conv_record env sigma (t1,sk1) (t2,sk2) = let params1, c1, extra_args1 = match arg with | Some c -> (* A primitive projection applied to c *) - let ty = Retyping.get_type_of ~lax:true env sigma c in + let ty = Retyping.get_type_of ~lax:true env sigma (EConstr.of_constr c) in let (i,u), ind_args = try Inductiveops.find_mrectype env sigma (EConstr.of_constr ty) with _ -> raise Not_found @@ -464,7 +464,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty else let f = try - let termM' = Retyping.expand_projection env evd p (EConstr.Unsafe.to_constr c) [] in + let termM' = Retyping.expand_projection env evd p c [] in let apprM', cstsM' = whd_betaiota_deltazeta_for_iota_state (fst ts) env evd cstsM (EConstr.of_constr termM',skM) @@ -643,7 +643,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty (* Catch the p.c ~= p c' cases *) | Proj (p,c), Const (p',u) when eq_constant (Projection.constant p) p' -> let res = - try Some (EConstr.destApp evd (EConstr.of_constr (Retyping.expand_projection env evd p c []))) + try Some (EConstr.destApp evd (EConstr.of_constr (Retyping.expand_projection env evd p (EConstr.of_constr c) []))) with Retyping.RetypeError _ -> None in (match res with @@ -654,7 +654,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty | Const (p,u), Proj (p',c') when eq_constant p (Projection.constant p') -> let res = - try Some (EConstr.destApp evd (EConstr.of_constr (Retyping.expand_projection env evd p' c' []))) + try Some (EConstr.destApp evd (EConstr.of_constr (Retyping.expand_projection env evd p' (EConstr.of_constr c') []))) with Retyping.RetypeError _ -> None in (match res with @@ -888,7 +888,7 @@ and conv_record trs env evd (ctx,(h,h2),c,bs,(params,params1),(us,us2),(sk1,sk2) List.fold_left (fun (i,ks,m,test) b -> if match n with Some n -> Int.equal m n | None -> false then - let ty = Retyping.get_type_of env i t2 in + let ty = Retyping.get_type_of env i (EConstr.of_constr t2) in let test i = evar_conv_x trs env i CUMUL ty (substl ks b) in (i,t2::ks, m-1, test) else @@ -1058,7 +1058,7 @@ let second_order_matching ts env_rhs evd (evk,args) argoccs rhs = let id = NamedDecl.get_id decl' in let t = NamedDecl.get_type decl' in let evs = ref [] in - let ty = Retyping.get_type_of env_rhs evd c in + let ty = Retyping.get_type_of env_rhs evd (EConstr.of_constr c) in let filter' = filter_possible_projections evd c ty ctxt args in (id,t,c,ty,evs,Filter.make filter',occs) :: make_subst (ctxt',l,occsl) | _, _, [] -> [] diff --git a/pretyping/evarsolve.ml b/pretyping/evarsolve.ml index 17bb1406e..86ef8f56f 100644 --- a/pretyping/evarsolve.ml +++ b/pretyping/evarsolve.ml @@ -141,8 +141,8 @@ let recheck_applications conv_algo env evdref t = match kind_of_term t with | App (f, args) -> let () = aux env f in - let fty = Retyping.get_type_of env !evdref f in - let argsty = Array.map (fun x -> aux env x; Retyping.get_type_of env !evdref x) args in + let fty = Retyping.get_type_of env !evdref (EConstr.of_constr f) in + let argsty = Array.map (fun x -> aux env x; Retyping.get_type_of env !evdref (EConstr.of_constr x)) args in let rec aux i ty = if i < Array.length argsty then match kind_of_term (whd_all env !evdref (EConstr.of_constr ty)) with @@ -634,7 +634,7 @@ let materialize_evar define_fun env evd k (evk1,args1) ty_in_env = let LocalAssum (na,t_in_env) | LocalDef (na,_,t_in_env) = d in let id = next_name_away na avoid in let evd,t_in_sign = - let s = Retyping.get_sort_of env evd t_in_env in + let s = Retyping.get_sort_of env evd (EConstr.of_constr t_in_env) in let evd,ty_t_in_sign = refresh_universes ~status:univ_flexible (Some false) env evd (mkSort s) in define_evar_from_virtual_equation define_fun env evd src t_in_env @@ -653,7 +653,7 @@ let materialize_evar define_fun env evd k (evk1,args1) ty_in_env = (sign1,filter1,Array.to_list args1,inst_in_sign,env1,evd,ids1) in let evd,ev2ty_in_sign = - let s = Retyping.get_sort_of env evd ty_in_env in + let s = Retyping.get_sort_of env evd (EConstr.of_constr ty_in_env) in let evd,ty_t_in_sign = refresh_universes ~status:univ_flexible (Some false) env evd (mkSort s) in define_evar_from_virtual_equation define_fun env evd src ty_in_env @@ -1161,7 +1161,7 @@ let check_evar_instance evd evk1 body conv_algo = (* FIXME: The body might be ill-typed when this is called from w_merge *) (* This happens in practice, cf MathClasses build failure on 2013-3-15 *) let ty = - try Retyping.get_type_of ~lax:true evenv evd body + try Retyping.get_type_of ~lax:true evenv evd (EConstr.of_constr body) with Retyping.RetypeError _ -> error "Ill-typed evar instance" in match conv_algo evenv evd Reduction.CUMUL ty evi.evar_concl with @@ -1365,7 +1365,7 @@ let rec invert_definition conv_algo choose env evd pbty (evk,argsv as ev) rhs = | (id,p)::_::_ -> if choose then (mkVar id, p) else raise (NotUniqueInType sols) in - let ty = lazy (Retyping.get_type_of env !evdref t) in + let ty = lazy (Retyping.get_type_of env !evdref (EConstr.of_constr t)) in let evd = do_projection_effects (evar_define conv_algo ~choose) env ty !evdref p in evdref := evd; c @@ -1428,7 +1428,7 @@ let rec invert_definition conv_algo choose env evd pbty (evk,argsv as ev) rhs = if not !progress then raise (NotEnoughInformationEvarEvar t); (* Make the virtual left evar real *) - let ty = get_type_of env' evd t in + let ty = get_type_of env' evd (EConstr.of_constr t) in let (evd,evar'',ev'') = materialize_evar (evar_define conv_algo ~choose) env' evd k ev ty in (* materialize_evar may instantiate ev' by another evar; adjust it *) @@ -1462,7 +1462,7 @@ let rec invert_definition conv_algo choose env evd pbty (evk,argsv as ev) rhs = | _ -> None with | Some l -> - let ty = get_type_of env' !evdref t in + let ty = get_type_of env' !evdref (EConstr.of_constr t) in let candidates = try let self env c = EConstr.of_constr (imitate env (EConstr.Unsafe.to_constr c)) in diff --git a/pretyping/evarsolve.mli b/pretyping/evarsolve.mli index 70e94b4dc..ac082d1bf 100644 --- a/pretyping/evarsolve.mli +++ b/pretyping/evarsolve.mli @@ -76,4 +76,4 @@ val remove_instance_local_defs : evar_map -> existential_key -> constr array -> constr list val get_type_of_refresh : - ?polyprop:bool -> ?lax:bool -> env -> evar_map -> constr -> evar_map * types + ?polyprop:bool -> ?lax:bool -> env -> evar_map -> EConstr.constr -> evar_map * types diff --git a/pretyping/inductiveops.ml b/pretyping/inductiveops.ml index a3cca2ad8..a9184777d 100644 --- a/pretyping/inductiveops.ml +++ b/pretyping/inductiveops.ml @@ -606,7 +606,7 @@ let rec instantiate_universes env evdref scl is = function let type_of_inductive_knowing_conclusion env sigma ((mib,mip),u) conclty = match mip.mind_arity with - | RegularArity s -> sigma, subst_instance_constr u s.mind_user_arity + | RegularArity s -> sigma, EConstr.of_constr (subst_instance_constr u s.mind_user_arity) | TemplateArity ar -> let _,scl = splay_arity env sigma conclty in let ctx = List.rev mip.mind_arity_ctxt in @@ -614,7 +614,7 @@ let type_of_inductive_knowing_conclusion env sigma ((mib,mip),u) conclty = let ctx = instantiate_universes env evdref scl ar.template_level (ctx,ar.template_param_levels) in - !evdref, mkArity (List.rev ctx,scl) + !evdref, EConstr.of_constr (mkArity (List.rev ctx,scl)) let type_of_projection_knowing_arg env sigma p c ty = let c = EConstr.Unsafe.to_constr c in diff --git a/pretyping/inductiveops.mli b/pretyping/inductiveops.mli index 1cfdef6e5..e375a2c6b 100644 --- a/pretyping/inductiveops.mli +++ b/pretyping/inductiveops.mli @@ -195,7 +195,7 @@ i*) (********************) val type_of_inductive_knowing_conclusion : - env -> evar_map -> Inductive.mind_specif puniverses -> EConstr.types -> evar_map * types + env -> evar_map -> Inductive.mind_specif puniverses -> EConstr.types -> evar_map * EConstr.types (********************) val control_only_guard : env -> types -> unit diff --git a/pretyping/patternops.ml b/pretyping/patternops.ml index 9dcb5d2a5..938b6b18e 100644 --- a/pretyping/patternops.ml +++ b/pretyping/patternops.ml @@ -153,7 +153,7 @@ let pattern_of_constr env sigma t = | Ind (sp,u) -> PRef (canonical_gr (IndRef sp)) | Construct (sp,u) -> PRef (canonical_gr (ConstructRef sp)) | Proj (p, c) -> - pattern_of_constr env (Retyping.expand_projection env sigma p c []) + pattern_of_constr env (Retyping.expand_projection env sigma p (EConstr.of_constr c) []) | Evar (evk,ctxt as ev) -> (match snd (Evd.evar_source evk sigma) with | Evar_kinds.MatchingVar (b,id) -> diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml index 2f42ad395..3c48c42df 100644 --- a/pretyping/pretyping.ml +++ b/pretyping/pretyping.ml @@ -409,7 +409,7 @@ let invert_ltac_bound_name lvar env id0 id = str " which is not bound in current context.") let protected_get_type_of env sigma c = - try Retyping.get_type_of ~lax:true env.ExtraEnv.env sigma c + try Retyping.get_type_of ~lax:true env.ExtraEnv.env sigma (EConstr.of_constr c) with Retyping.RetypeError _ -> user_err (str "Cannot reinterpret " ++ quote (print_constr c) ++ @@ -563,7 +563,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre let hyps = evar_filtered_context (Evd.find !evdref evk) in let args = pretype_instance k0 resolve_tc env evdref lvar loc hyps evk inst in let c = mkEvar (evk, args) in - let j = (Retyping.get_judgment_of env.ExtraEnv.env !evdref c) in + let j = (Retyping.get_judgment_of env.ExtraEnv.env !evdref (EConstr.of_constr c)) in inh_conv_coerce_to_tycon loc env evdref j tycon | GPatVar (loc,(someta,n)) -> @@ -757,7 +757,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre let sigma = !evdref in let c = mkApp (f,Array.map (whd_evar sigma) args) in let c = evd_comb1 (Evarsolve.refresh_universes (Some true) env.ExtraEnv.env) evdref c in - let t = Retyping.get_type_of env.ExtraEnv.env !evdref c in + let t = Retyping.get_type_of env.ExtraEnv.env !evdref (EConstr.of_constr c) in make_judge c (* use this for keeping evars: resj.uj_val *) t else resj | _ -> resj @@ -1067,7 +1067,7 @@ and pretype_type k0 resolve_tc valcon (env : ExtraEnv.t) evdref lvar = function | Some v -> let s = let sigma = !evdref in - let t = Retyping.get_type_of env.ExtraEnv.env sigma v in + let t = Retyping.get_type_of env.ExtraEnv.env sigma (EConstr.of_constr v) in match kind_of_term (whd_all env.ExtraEnv.env sigma (EConstr.of_constr t)) with | Sort s -> s | Evar ev when is_Type (existential_type sigma ev) -> diff --git a/pretyping/retyping.ml b/pretyping/retyping.ml index 353bdbb89..2efb02417 100644 --- a/pretyping/retyping.ml +++ b/pretyping/retyping.ml @@ -49,12 +49,21 @@ let anomaly_on_error f x = try f x with RetypeError e -> anomaly ~label:"retyping" (print_retype_error e) +let local_assum (na, t) = + let inj = EConstr.Unsafe.to_constr in + LocalAssum (na, inj t) + +let local_def (na, b, t) = + let inj = EConstr.Unsafe.to_constr in + LocalDef (na, inj b, inj t) + let get_type_from_constraints env sigma t = - if isEvar (fst (decompose_app t)) then + let open EConstr in + if isEvar sigma (EConstr.of_constr (fst (decompose_app_vect sigma t))) then match List.map_filter (fun (pbty,env,t1,t2) -> - if is_fconv Reduction.CONV env sigma (EConstr.of_constr t) (EConstr.of_constr t1) then Some t2 - else if is_fconv Reduction.CONV env sigma (EConstr.of_constr t) (EConstr.of_constr t2) then Some t1 + if is_fconv Reduction.CONV env sigma t (EConstr.of_constr t1) then Some t2 + else if is_fconv Reduction.CONV env sigma t (EConstr.of_constr t2) then Some t1 else None) (snd (Evd.extract_all_conv_pbs sigma)) with @@ -65,87 +74,89 @@ let get_type_from_constraints env sigma t = let rec subst_type env sigma typ = function | [] -> typ | h::rest -> - match kind_of_term (whd_all env sigma (EConstr.of_constr typ)) with - | Prod (na,c1,c2) -> subst_type env sigma (subst1 h c2) rest + let open EConstr in + match EConstr.kind sigma (EConstr.of_constr (whd_all env sigma typ)) with + | Prod (na,c1,c2) -> subst_type env sigma (Vars.subst1 h c2) rest | _ -> retype_error NonFunctionalConstruction (* If ft is the type of f which itself is applied to args, *) (* [sort_of_atomic_type] computes ft[args] which has to be a sort *) let sort_of_atomic_type env sigma ft args = + let open EConstr in let rec concl_of_arity env n ar args = - match kind_of_term (whd_all env sigma (EConstr.of_constr ar)), args with - | Prod (na, t, b), h::l -> concl_of_arity (push_rel (LocalDef (na, lift n h, t)) env) (n + 1) b l + match EConstr.kind sigma (EConstr.of_constr (whd_all env sigma ar)), args with + | Prod (na, t, b), h::l -> concl_of_arity (push_rel (local_def (na, Vars.lift n h, t)) env) (n + 1) b l | Sort s, [] -> s | _ -> retype_error NotASort in concl_of_arity env 0 ft (Array.to_list args) let type_of_var env id = - try NamedDecl.get_type (lookup_named id env) + try EConstr.of_constr (NamedDecl.get_type (lookup_named id env)) with Not_found -> retype_error (BadVariable id) let decomp_sort env sigma t = - match kind_of_term (whd_all env sigma t) with + match EConstr.kind sigma (EConstr.of_constr (whd_all env sigma t)) with | Sort s -> s | _ -> retype_error NotASort let retype ?(polyprop=true) sigma = + let open EConstr in let rec type_of env cstr = - match kind_of_term cstr with + match EConstr.kind sigma cstr with | Meta n -> - (try strip_outer_cast sigma (EConstr.of_constr (Evd.meta_ftype sigma n).Evd.rebus) + EConstr.of_constr (try strip_outer_cast sigma (EConstr.of_constr (Evd.meta_ftype sigma n).Evd.rebus) with Not_found -> retype_error (BadMeta n)) | Rel n -> - let ty = RelDecl.get_type (lookup_rel n env) in - lift n ty + let ty = EConstr.of_constr (RelDecl.get_type (lookup_rel n env)) in + Vars.lift n ty | Var id -> type_of_var env id - | Const cst -> rename_type_of_constant env cst - | Evar ev -> Evd.existential_type sigma ev - | Ind ind -> rename_type_of_inductive env ind - | Construct cstr -> rename_type_of_constructor env cstr + | Const cst -> EConstr.of_constr (rename_type_of_constant env cst) + | Evar (evk, args) -> EConstr.of_constr (Evd.existential_type sigma (evk, Array.map EConstr.Unsafe.to_constr args)) + | Ind ind -> EConstr.of_constr (rename_type_of_inductive env ind) + | Construct cstr -> EConstr.of_constr (rename_type_of_constructor env cstr) | Case (_,p,c,lf) -> let Inductiveops.IndType(indf,realargs) = let t = type_of env c in - try Inductiveops.find_rectype env sigma (EConstr.of_constr t) + try Inductiveops.find_rectype env sigma t with Not_found -> try - let t = get_type_from_constraints env sigma t in - Inductiveops.find_rectype env sigma (EConstr.of_constr t) + let t = EConstr.of_constr (get_type_from_constraints env sigma t) in + Inductiveops.find_rectype env sigma t with Not_found -> retype_error BadRecursiveType in let n = inductive_nrealdecls_env env (fst (fst (dest_ind_family indf))) in - let t = betazetaevar_applist sigma n (EConstr.of_constr p) (List.map EConstr.of_constr realargs) in - (match kind_of_term (whd_all env sigma (EConstr.of_constr (type_of env t))) with - | Prod _ -> whd_beta sigma (EConstr.of_constr (applist (t, [c]))) + let t = EConstr.of_constr (betazetaevar_applist sigma n p (List.map EConstr.of_constr realargs)) in + (match EConstr.kind sigma (EConstr.of_constr (whd_all env sigma (type_of env t))) with + | Prod _ -> EConstr.of_constr (whd_beta sigma (applist (t, [c]))) | _ -> t) | Lambda (name,c1,c2) -> - mkProd (name, c1, type_of (push_rel (LocalAssum (name,c1)) env) c2) + mkProd (name, c1, type_of (push_rel (local_assum (name,c1)) env) c2) | LetIn (name,b,c1,c2) -> - subst1 b (type_of (push_rel (LocalDef (name,b,c1)) env) c2) + Vars.subst1 b (type_of (push_rel (local_def (name,b,c1)) env) c2) | Fix ((_,i),(_,tys,_)) -> tys.(i) | CoFix (i,(_,tys,_)) -> tys.(i) - | App(f,args) when is_template_polymorphic env sigma (EConstr.of_constr f) -> - let f = whd_evar sigma f in + | App(f,args) when is_template_polymorphic env sigma f -> let t = type_of_global_reference_knowing_parameters env f args in - strip_outer_cast sigma (EConstr.of_constr (subst_type env sigma t (Array.to_list args))) + EConstr.of_constr (strip_outer_cast sigma (subst_type env sigma t (Array.to_list args))) | App(f,args) -> - strip_outer_cast sigma - (EConstr.of_constr (subst_type env sigma (type_of env f) (Array.to_list args))) + EConstr.of_constr (strip_outer_cast sigma + (subst_type env sigma (type_of env f) (Array.to_list args))) | Proj (p,c) -> let ty = type_of env c in - (try - Inductiveops.type_of_projection_knowing_arg env sigma p (EConstr.of_constr c) (EConstr.of_constr ty) + EConstr.of_constr (try + Inductiveops.type_of_projection_knowing_arg env sigma p c ty with Invalid_argument _ -> retype_error BadRecursiveType) | Cast (c,_, t) -> t | Sort _ | Prod _ -> mkSort (sort_of env cstr) and sort_of env t = - match kind_of_term t with - | Cast (c,_, s) when isSort s -> destSort s + match EConstr.kind sigma t with + | Cast (c,_, s) when isSort sigma s -> destSort sigma s | Sort (Prop c) -> type1_sort | Sort (Type u) -> Type (Univ.super u) | Prod (name,t,c2) -> - (match (sort_of env t, sort_of (push_rel (LocalAssum (name,t)) env) c2) with + (match (sort_of env t, sort_of (push_rel (local_assum (name,t)) env) c2) with | _, (Prop Null as s) -> s | Prop _, (Prop Pos as s) -> s | Type _, (Prop Pos as s) when is_impredicative_set env -> s @@ -153,47 +164,45 @@ let retype ?(polyprop=true) sigma = | Prop Pos, (Type u2) -> Type (Univ.sup Univ.type0_univ u2) | Prop Null, (Type _ as s) -> s | Type u1, Type u2 -> Type (Univ.sup u1 u2)) - | App(f,args) when is_template_polymorphic env sigma (EConstr.of_constr f) -> - let f = whd_evar sigma f in + | App(f,args) when is_template_polymorphic env sigma f -> let t = type_of_global_reference_knowing_parameters env f args in sort_of_atomic_type env sigma t args | App(f,args) -> sort_of_atomic_type env sigma (type_of env f) args | Lambda _ | Fix _ | Construct _ -> retype_error NotAType - | _ -> decomp_sort env sigma (EConstr.of_constr (type_of env t)) + | _ -> decomp_sort env sigma (type_of env t) and sort_family_of env t = - match kind_of_term t with - | Cast (c,_, s) when isSort s -> family_of_sort (destSort s) + match EConstr.kind sigma t with + | Cast (c,_, s) when isSort sigma s -> family_of_sort (destSort sigma s) | Sort (Prop c) -> InType | Sort (Type u) -> InType | Prod (name,t,c2) -> - let s2 = sort_family_of (push_rel (LocalAssum (name,t)) env) c2 in + let s2 = sort_family_of (push_rel (local_assum (name,t)) env) c2 in if not (is_impredicative_set env) && s2 == InSet && sort_family_of env t == InType then InType else s2 - | App(f,args) when is_template_polymorphic env sigma (EConstr.of_constr f) -> - let f = whd_evar sigma f in + | App(f,args) when is_template_polymorphic env sigma f -> let t = type_of_global_reference_knowing_parameters env f args in family_of_sort (sort_of_atomic_type env sigma t args) | App(f,args) -> family_of_sort (sort_of_atomic_type env sigma (type_of env f) args) | Lambda _ | Fix _ | Construct _ -> retype_error NotAType | _ -> - family_of_sort (decomp_sort env sigma (EConstr.of_constr (type_of env t))) + family_of_sort (decomp_sort env sigma (type_of env t)) and type_of_global_reference_knowing_parameters env c args = let argtyps = - Array.map (fun c -> lazy (nf_evar sigma (type_of env c))) args in - match kind_of_term c with + Array.map (fun c -> lazy (EConstr.to_constr sigma (type_of env c))) args in + match EConstr.kind sigma c with | Ind ind -> let mip = lookup_mind_specif env (fst ind) in - (try Inductive.type_of_inductive_knowing_parameters + EConstr.of_constr (try Inductive.type_of_inductive_knowing_parameters ~polyprop env (mip,snd ind) argtyps with Reduction.NotArity -> retype_error NotAnArity) | Const cst -> - (try Typeops.type_of_constant_knowing_parameters_in env cst argtyps + EConstr.of_constr (try Typeops.type_of_constant_knowing_parameters_in env cst argtyps with Reduction.NotArity -> retype_error NotAnArity) | Var id -> type_of_var env id - | Construct cstr -> type_of_constructor env cstr + | Construct cstr -> EConstr.of_constr (type_of_constructor env cstr) | _ -> assert false in type_of, sort_of, sort_family_of, @@ -204,19 +213,19 @@ let get_sort_of ?(polyprop=true) env sigma t = let get_sort_family_of ?(polyprop=true) env sigma c = let _,_,f,_ = retype ~polyprop sigma in anomaly_on_error (f env) c let type_of_global_reference_knowing_parameters env sigma c args = - let _,_,_,f = retype sigma in anomaly_on_error (f env c) args + let _,_,_,f = retype sigma in EConstr.Unsafe.to_constr (anomaly_on_error (f env c) args) let type_of_global_reference_knowing_conclusion env sigma c conclty = - match kind_of_term c with + match EConstr.kind sigma c with | Ind (ind,u) -> let spec = Inductive.lookup_mind_specif env ind in - type_of_inductive_knowing_conclusion env sigma (spec,u) (EConstr.of_constr conclty) + type_of_inductive_knowing_conclusion env sigma (spec,u) conclty | Const cst -> let t = constant_type_in env cst in (* TODO *) - sigma, Typeops.type_of_constant_type_knowing_parameters env t [||] + sigma, EConstr.of_constr (Typeops.type_of_constant_type_knowing_parameters env t [||]) | Var id -> sigma, type_of_var env id - | Construct cstr -> sigma, type_of_constructor env cstr + | Construct cstr -> sigma, EConstr.of_constr (type_of_constructor env cstr) | _ -> assert false (* Profiling *) @@ -232,10 +241,10 @@ let type_of_global_reference_knowing_conclusion env sigma c conclty = let get_type_of ?(polyprop=true) ?(lax=false) env sigma c = let f,_,_,_ = retype ~polyprop sigma in - if lax then f env c else anomaly_on_error (f env) c + if lax then EConstr.Unsafe.to_constr (f env c) else EConstr.Unsafe.to_constr (anomaly_on_error (f env) c) (* Makes an unsafe judgment from a constr *) -let get_judgment_of env evc c = { uj_val = c; uj_type = get_type_of env evc c } +let get_judgment_of env evc c = { uj_val = EConstr.Unsafe.to_constr c; uj_type = get_type_of env evc c } (* Returns sorts of a context *) let sorts_of_context env evc ctxt = @@ -243,15 +252,16 @@ let sorts_of_context env evc ctxt = | [] -> env,[] | d :: ctxt -> let env,sorts = aux ctxt in - let s = get_sort_of env evc (RelDecl.get_type d) in + let s = get_sort_of env evc (EConstr.of_constr (RelDecl.get_type d)) in (push_rel d env,s::sorts) in snd (aux ctxt) let expand_projection env sigma pr c args = + let inj = EConstr.Unsafe.to_constr in let ty = get_type_of ~lax:true env sigma c in let (i,u), ind_args = try Inductiveops.find_mrectype env sigma (EConstr.of_constr ty) with Not_found -> retype_error BadRecursiveType in mkApp (mkConstU (Projection.constant pr,u), - Array.of_list (ind_args @ (c :: args))) + Array.of_list (ind_args @ (inj c :: List.map inj args))) diff --git a/pretyping/retyping.mli b/pretyping/retyping.mli index 8ca40f829..08f750287 100644 --- a/pretyping/retyping.mli +++ b/pretyping/retyping.mli @@ -26,25 +26,25 @@ type retype_error exception RetypeError of retype_error val get_type_of : - ?polyprop:bool -> ?lax:bool -> env -> evar_map -> constr -> types + ?polyprop:bool -> ?lax:bool -> env -> evar_map -> EConstr.constr -> types val get_sort_of : - ?polyprop:bool -> env -> evar_map -> types -> sorts + ?polyprop:bool -> env -> evar_map -> EConstr.types -> sorts val get_sort_family_of : - ?polyprop:bool -> env -> evar_map -> types -> sorts_family + ?polyprop:bool -> env -> evar_map -> EConstr.types -> sorts_family (** Makes an unsafe judgment from a constr *) -val get_judgment_of : env -> evar_map -> constr -> unsafe_judgment +val get_judgment_of : env -> evar_map -> EConstr.constr -> unsafe_judgment -val type_of_global_reference_knowing_parameters : env -> evar_map -> constr -> - constr array -> types +val type_of_global_reference_knowing_parameters : env -> evar_map -> EConstr.constr -> + EConstr.constr array -> types val type_of_global_reference_knowing_conclusion : - env -> evar_map -> constr -> types -> evar_map * types + env -> evar_map -> EConstr.constr -> EConstr.types -> evar_map * EConstr.types val sorts_of_context : env -> evar_map -> Context.Rel.t -> sorts list -val expand_projection : env -> evar_map -> Names.projection -> constr -> constr list -> constr +val expand_projection : env -> evar_map -> Names.projection -> EConstr.constr -> EConstr.constr list -> constr val print_retype_error : retype_error -> Pp.std_ppcmds diff --git a/pretyping/tacred.ml b/pretyping/tacred.ml index 357a80f44..ac2a3bc49 100644 --- a/pretyping/tacred.ml +++ b/pretyping/tacred.ml @@ -974,7 +974,7 @@ let matches_head env sigma c t = let change_map_constr_with_binders_left_to_right g f (env, l as acc) sigma c = match kind_of_term c with | Proj (p, r) -> (* Treat specially for partial applications *) - let t = Retyping.expand_projection env sigma p r [] in + let t = Retyping.expand_projection env sigma p (EConstr.of_constr r) [] in let hdf, al = destApp t in let a = al.(Array.length al - 1) in let app = (mkApp (hdf, Array.sub al 0 (Array.length al - 1))) in @@ -1150,7 +1150,7 @@ let compute = cbv_betadeltaiota * the specified occurrences. *) let abstract_scheme env (locc,a) (c, sigma) = - let ta = Retyping.get_type_of env sigma a in + let ta = Retyping.get_type_of env sigma (EConstr.of_constr a) in let na = named_hd env ta Anonymous in if occur_meta sigma (EConstr.of_constr ta) then error "Cannot find a type for the generalisation."; if occur_meta sigma (EConstr.of_constr a) then diff --git a/pretyping/typeclasses.ml b/pretyping/typeclasses.ml index 50dd66ea0..8c03329e2 100644 --- a/pretyping/typeclasses.ml +++ b/pretyping/typeclasses.ml @@ -300,7 +300,7 @@ let build_subclasses ~check env sigma glob pri = | Some (Forward, pri') -> let proj = Option.get proj in let body = it_mkLambda_or_LetIn (mkApp (mkConstU (proj,u), projargs)) rels in - if check && check_instance env sigma body then None + if check && check_instance env sigma (EConstr.of_constr body) then None else let pri = match pri, pri' with @@ -312,7 +312,7 @@ let build_subclasses ~check env sigma glob pri = in let declare_proj hints (cref, pri, body) = let path' = cref :: path in - let ty = Retyping.get_type_of env sigma body in + let ty = Retyping.get_type_of env sigma (EConstr.of_constr body) in let rest = aux pri body ty path' in hints @ (path', pri, body) :: rest in List.fold_left declare_proj [] projs diff --git a/pretyping/typing.ml b/pretyping/typing.ml index e3d1c1741..acfe05f24 100644 --- a/pretyping/typing.ml +++ b/pretyping/typing.ml @@ -138,7 +138,7 @@ let check_type_fixpoint loc env evdref lna lar vdefj = (* FIXME: might depend on the level of actual parameters!*) let check_allowed_sort env sigma ind c p = - let pj = Retyping.get_judgment_of env sigma p in + let pj = Retyping.get_judgment_of env sigma (EConstr.of_constr p) in let ksort = family_of_sort (sort_of_arity env sigma (EConstr.of_constr pj.uj_type)) in let specif = Global.lookup_inductive (fst ind) in let sorts = elim_sorts specif in diff --git a/pretyping/unification.ml b/pretyping/unification.ml index ede2606da..848a2f4a8 100644 --- a/pretyping/unification.ml +++ b/pretyping/unification.ml @@ -684,7 +684,7 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb if opt.with_types && flags.check_applied_meta_types then (try let tyM = Typing.meta_type sigma k in - let tyN = get_type_of curenv ~lax:true sigma cN in + let tyN = get_type_of curenv ~lax:true sigma (EConstr.of_constr cN) in check_compatibility curenv CUMUL flags substn tyN tyM with RetypeError _ -> (* Renounce, maybe metas/evars prevents typing *) sigma) @@ -703,7 +703,7 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb let sigma = if opt.with_types && flags.check_applied_meta_types then (try - let tyM = get_type_of curenv ~lax:true sigma cM in + let tyM = get_type_of curenv ~lax:true sigma (EConstr.of_constr cM) in let tyN = Typing.meta_type sigma k in check_compatibility curenv CUMUL flags substn tyM tyN with RetypeError _ -> @@ -863,7 +863,7 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb let expand_proj c c' l = match kind_of_term c with | Proj (p, t) when not (Projection.unfolded p) && needs_expansion p c' -> - (try destApp (Retyping.expand_projection curenv sigma p t (Array.to_list l)) + (try destApp (Retyping.expand_projection curenv sigma p (EConstr.of_constr t) (Array.map_to_list EConstr.of_constr l)) with RetypeError _ -> (** Unification can be called on ill-typed terms, due to FO and eta in particular, fail gracefully in that case *) (c, l)) @@ -888,8 +888,8 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb and unify_same_proj (curenv, nb as curenvnb) cv_pb opt substn c1 c2 = let substn = unirec_rec curenvnb CONV opt substn c1 c2 in try (* Force unification of the types to fill in parameters *) - let ty1 = get_type_of curenv ~lax:true sigma c1 in - let ty2 = get_type_of curenv ~lax:true sigma c2 in + let ty1 = get_type_of curenv ~lax:true sigma (EConstr.of_constr c1) in + let ty2 = get_type_of curenv ~lax:true sigma (EConstr.of_constr c2) in unify_0_with_initial_metas substn true curenv cv_pb { flags with modulo_conv_on_closed_terms = Some full_transparent_state; modulo_delta = full_transparent_state; @@ -953,8 +953,8 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb let sigma = if opt.with_types then try (* Ensure we call conversion on terms of the same type *) - let tyM = get_type_of curenv ~lax:true sigma m1 in - let tyN = get_type_of curenv ~lax:true sigma n1 in + let tyM = get_type_of curenv ~lax:true sigma (EConstr.of_constr m1) in + let tyN = get_type_of curenv ~lax:true sigma (EConstr.of_constr n1) in check_compatibility curenv CUMUL flags substn tyM tyN with RetypeError _ -> (* Renounce, maybe metas/evars prevents typing *) sigma @@ -1240,13 +1240,13 @@ let w_coerce_to_type env evd c cty mvty = try_to_coerce env evd c cty tycon let w_coerce env evd mv c = - let cty = get_type_of env evd c in + let cty = get_type_of env evd (EConstr.of_constr c) in let mvty = Typing.meta_type evd mv in w_coerce_to_type env evd c cty mvty let unify_to_type env sigma flags c status u = let sigma, c = refresh_universes (Some false) env sigma c in - let t = get_type_of env sigma (nf_meta sigma c) in + let t = get_type_of env sigma (EConstr.of_constr (nf_meta sigma c)) in let t = nf_betaiota sigma (EConstr.of_constr (nf_meta sigma t)) in unify_0 env sigma CUMUL flags t u @@ -1379,7 +1379,7 @@ let w_merge env with_types flags (evd,metas,evars) = let evd' = Sigma.to_evar_map evd' in let (evd'',mc,ec) = unify_0 sp_env evd' CUMUL flags - (get_type_of sp_env evd' c) ev.evar_concl in + (get_type_of sp_env evd' (EConstr.of_constr c)) ev.evar_concl in let evd''' = w_merge_rec evd'' mc ec [] in if evd' == evd''' then Evd.define sp c evd''' @@ -1422,13 +1422,13 @@ let check_types env flags (sigma,_,_ as subst) m n = if isEvar_or_Meta (head_app sigma m) then unify_0_with_initial_metas subst true env CUMUL flags - (get_type_of env sigma n) - (get_type_of env sigma m) + (get_type_of env sigma (EConstr.of_constr n)) + (get_type_of env sigma (EConstr.of_constr m)) else if isEvar_or_Meta (head_app sigma n) then unify_0_with_initial_metas subst true env CUMUL flags - (get_type_of env sigma m) - (get_type_of env sigma n) + (get_type_of env sigma (EConstr.of_constr m)) + (get_type_of env sigma (EConstr.of_constr n)) else subst let try_resolve_typeclasses env evd flag m n = @@ -1557,7 +1557,7 @@ let make_pattern_test from_prefix_of_ind is_correct_type env sigma (pending,c) = applist (t,l1), l2 else t, [] in let sigma = w_typed_unify env sigma Reduction.CONV flags c t' in - let ty = Retyping.get_type_of env sigma t in + let ty = Retyping.get_type_of env sigma (EConstr.of_constr t) in if not (is_correct_type ty) then raise (NotUnifiable None); Some(sigma, t, l2) with @@ -1590,7 +1590,7 @@ let make_eq_test env evd c = let make_abstraction_core name (test,out) env sigma c ty occs check_occs concl = let id = - let t = match ty with Some t -> t | None -> get_type_of env sigma c in + let t = match ty with Some t -> t | None -> get_type_of env sigma (EConstr.of_constr c) in let x = id_of_name_using_hdchar (Global.env()) t name in let ids = ids_of_named_context (named_context env) in if name == Anonymous then next_ident_away_in_goal x ids else diff --git a/proofs/clenv.ml b/proofs/clenv.ml index d64cd9fff..f4ac094b8 100644 --- a/proofs/clenv.ml +++ b/proofs/clenv.ml @@ -67,7 +67,7 @@ let refresh_undefined_univs clenv = let clenv_hnf_constr ce t = hnf_constr (cl_env ce) (cl_sigma ce) t -let clenv_get_type_of ce c = Retyping.get_type_of (cl_env ce) (cl_sigma ce) c +let clenv_get_type_of ce c = Retyping.get_type_of (cl_env ce) (cl_sigma ce) (EConstr.of_constr c) exception NotExtensibleClause @@ -667,8 +667,8 @@ let evar_of_binder holes = function user_err (str "No such binder.") let define_with_type sigma env ev c = - let t = Retyping.get_type_of env sigma ev in - let ty = Retyping.get_type_of env sigma c in + let t = Retyping.get_type_of env sigma (EConstr.of_constr ev) in + let ty = Retyping.get_type_of env sigma (EConstr.of_constr c) in let j = Environ.make_judge c ty in let (sigma, j) = Coercion.inh_conv_coerce_to true (Loc.ghost) env sigma j t in let (ev, _) = destEvar ev in diff --git a/proofs/logic.ml b/proofs/logic.ml index 0fa193065..2df626e1c 100644 --- a/proofs/logic.ml +++ b/proofs/logic.ml @@ -319,18 +319,19 @@ let check_conv_leq_goal env sigma arg ty conclty = else raise (RefinerError (BadType (arg,ty,conclty))) else sigma -exception Stop of constr list +exception Stop of EConstr.t list let meta_free_prefix sigma a = try + let a = Array.map EConstr.of_constr a in let _ = Array.fold_left (fun acc a -> - if occur_meta sigma (EConstr.of_constr a) then raise (Stop acc) + if occur_meta sigma a then raise (Stop acc) else a :: acc) [] a in a with Stop acc -> Array.rev_of_list acc let goal_type_of env sigma c = if !check then unsafe_type_of env sigma c - else Retyping.get_type_of env sigma c + else Retyping.get_type_of env sigma (EConstr.of_constr c) let rec mk_refgoals sigma goal goalacc conclty trm = let env = Goal.V82.env sigma goal in @@ -339,7 +340,7 @@ let rec mk_refgoals sigma goal goalacc conclty trm = Goal.V82.mk_goal sigma hyps concl (Goal.V82.extra sigma goal) in if (not !check) && not (occur_meta sigma (EConstr.of_constr trm)) then - let t'ty = Retyping.get_type_of env sigma trm in + let t'ty = Retyping.get_type_of env sigma (EConstr.of_constr trm) in let sigma = check_conv_leq_goal env sigma trm t'ty conclty in (goalacc,t'ty,sigma,trm) else @@ -372,7 +373,7 @@ let rec mk_refgoals sigma goal goalacc conclty trm = (* Template sort-polymorphism of definition and inductive types *) let firstmeta = Array.findi (fun i x -> occur_meta sigma (EConstr.of_constr x)) l in let args, _ = Option.cata (fun i -> CArray.chop i l) (l, [||]) firstmeta in - type_of_global_reference_knowing_parameters env sigma f args + type_of_global_reference_knowing_parameters env sigma (EConstr.of_constr f) (Array.map EConstr.of_constr args) in goalacc, ty, sigma, f else @@ -386,7 +387,7 @@ let rec mk_refgoals sigma goal goalacc conclty trm = | Proj (p,c) -> let (acc',cty,sigma,c') = mk_hdgoals sigma goal goalacc c in let c = mkProj (p, c') in - let ty = get_type_of env sigma c in + let ty = get_type_of env sigma (EConstr.of_constr c) in (acc',ty,sigma,c) | Case (ci,p,c,lf) -> @@ -435,7 +436,7 @@ and mk_hdgoals sigma goal goalacc trm = if is_template_polymorphic env sigma (EConstr.of_constr f) then let l' = meta_free_prefix sigma l in - (goalacc,type_of_global_reference_knowing_parameters env sigma f l',sigma,f) + (goalacc,type_of_global_reference_knowing_parameters env sigma (EConstr.of_constr f) l',sigma,f) else mk_hdgoals sigma goal goalacc f in let ((acc'',conclty',sigma), args) = mk_arggoals sigma goal acc' hdty l in @@ -460,7 +461,7 @@ and mk_hdgoals sigma goal goalacc trm = | Proj (p,c) -> let (acc',cty,sigma,c') = mk_hdgoals sigma goal goalacc c in let c = mkProj (p, c') in - let ty = get_type_of env sigma c in + let ty = get_type_of env sigma (EConstr.of_constr c) in (acc',ty,sigma,c) | _ -> diff --git a/proofs/redexpr.ml b/proofs/redexpr.ml index 62fe2c17c..19e72e697 100644 --- a/proofs/redexpr.ml +++ b/proofs/redexpr.ml @@ -24,7 +24,7 @@ open Misctypes (* call by value normalisation function using the virtual machine *) let cbv_vm env sigma c = - let ctyp = Retyping.get_type_of env sigma (EConstr.Unsafe.to_constr c) in + let ctyp = Retyping.get_type_of env sigma c in Vnorm.cbv_vm env sigma c (EConstr.of_constr ctyp) let warn_native_compute_disabled = @@ -37,7 +37,7 @@ let cbv_native env sigma c = (warn_native_compute_disabled (); cbv_vm env sigma c) else - let ctyp = Retyping.get_type_of env sigma (EConstr.Unsafe.to_constr c) in + let ctyp = Retyping.get_type_of env sigma c in Nativenorm.native_norm env sigma c (EConstr.of_constr ctyp) let whd_cbn flags env sigma t = diff --git a/proofs/refine.ml b/proofs/refine.ml index d4920e505..e6e3ef47d 100644 --- a/proofs/refine.ml +++ b/proofs/refine.ml @@ -133,7 +133,7 @@ let refine ?(unsafe = true) f = (** Useful definitions *) let with_type env evd c t = - let my_type = Retyping.get_type_of env evd c in + let my_type = Retyping.get_type_of env evd (EConstr.of_constr c) in let j = Environ.make_judge c my_type in let (evd,j') = Coercion.inh_conv_coerce_to true (Loc.ghost) env evd j t diff --git a/proofs/tacmach.ml b/proofs/tacmach.ml index 09f5246ab..b63b2ce28 100644 --- a/proofs/tacmach.ml +++ b/proofs/tacmach.ml @@ -104,7 +104,7 @@ let pf_const_value = pf_reduce (fun env _ -> constant_value_in env) let pf_reduce_to_quantified_ind = pf_reduce reduce_to_quantified_ind let pf_reduce_to_atomic_ind = pf_reduce reduce_to_atomic_ind -let pf_hnf_type_of gls = pf_get_type_of gls %> pf_whd_all gls +let pf_hnf_type_of gls = EConstr.of_constr %> pf_get_type_of gls %> pf_whd_all gls let pf_is_matching = pf_apply Constr_matching.is_matching_conv let pf_matches = pf_apply Constr_matching.matches_conv @@ -230,7 +230,7 @@ module New = struct let pf_hnf_constr gl t = pf_reduce hnf_constr gl t let pf_hnf_type_of gl t = - pf_whd_all gl (pf_get_type_of gl t) + pf_whd_all gl (pf_get_type_of gl (EConstr.of_constr t)) let pf_matches gl pat t = pf_apply Constr_matching.matches_conv gl pat t diff --git a/tactics/class_tactics.ml b/tactics/class_tactics.ml index fe7a09f77..6fb90e7af 100644 --- a/tactics/class_tactics.ml +++ b/tactics/class_tactics.ml @@ -280,7 +280,7 @@ let clenv_of_prods poly nprods (c, clenv) gl = if poly || Int.equal nprods 0 then Some (None, clenv) else let sigma = Tacmach.New.project gl in - let ty = Retyping.get_type_of (Proofview.Goal.env gl) sigma c in + let ty = Retyping.get_type_of (Proofview.Goal.env gl) sigma (EConstr.of_constr c) in let diff = nb_prod sigma (EConstr.of_constr ty) - nprods in if Pervasives.(>=) diff 0 then (* Was Some clenv... *) @@ -473,7 +473,7 @@ let catchable = function let pr_depth l = prlist_with_sep (fun () -> str ".") int (List.rev l) let is_Prop env sigma concl = - let ty = Retyping.get_type_of env sigma concl in + let ty = Retyping.get_type_of env sigma (EConstr.of_constr concl) in match kind_of_term ty with | Sort (Prop Null) -> true | _ -> false diff --git a/tactics/contradiction.ml b/tactics/contradiction.ml index b9704b846..789028ac1 100644 --- a/tactics/contradiction.ml +++ b/tactics/contradiction.ml @@ -28,7 +28,7 @@ let absurd c = let sigma = Proofview.Goal.sigma gl in let env = Proofview.Goal.env gl in let sigma = Sigma.to_evar_map sigma in - let j = Retyping.get_judgment_of env sigma c in + let j = Retyping.get_judgment_of env sigma (EConstr.of_constr c) in let sigma, j = Coercion.inh_coerce_to_sort Loc.ghost env sigma j in let t = j.Environ.utj_val in let tac = diff --git a/tactics/eqschemes.ml b/tactics/eqschemes.ml index aea3ca17e..92480e253 100644 --- a/tactics/eqschemes.ml +++ b/tactics/eqschemes.ml @@ -596,7 +596,7 @@ let build_r2l_forward_rew_scheme dep env ind kind = (**********************************************************************) let fix_r2l_forward_rew_scheme (c, ctx') = - let t = Retyping.get_type_of (Global.env()) Evd.empty c in + let t = Retyping.get_type_of (Global.env()) Evd.empty (EConstr.of_constr c) in let ctx,_ = decompose_prod_assum t in match ctx with | hp :: p :: ind :: indargs -> diff --git a/tactics/equality.ml b/tactics/equality.ml index 48f46b36b..e87746a28 100644 --- a/tactics/equality.ml +++ b/tactics/equality.ml @@ -440,7 +440,7 @@ let general_rewrite_ebindings_clause cls lft2rgt occs frzevars dep_proof_ok ?tac Proofview.Goal.enter { enter = begin fun gl -> let sigma = Tacmach.New.project gl in let env = Proofview.Goal.env gl in - let ctype = get_type_of env sigma c in + let ctype = get_type_of env sigma (EConstr.of_constr c) in let rels, t = decompose_prod_assum (whd_betaiotazeta sigma (EConstr.of_constr ctype)) in match match_with_equality_type sigma t with | Some (hdcncl,args) -> (* Fast path: direct leibniz-like rewrite *) @@ -621,8 +621,8 @@ let replace_using_leibniz clause c1 c2 l2r unsafe try_prove_eq_opt = in Proofview.Goal.enter { enter = begin fun gl -> let get_type_of = pf_apply get_type_of gl in - let t1 = get_type_of c1 - and t2 = get_type_of c2 in + let t1 = get_type_of (EConstr.of_constr c1) + and t2 = get_type_of (EConstr.of_constr c2) in let evd = if unsafe then Some (Tacmach.New.project gl) else @@ -719,8 +719,8 @@ let find_positions env sigma t1 t2 = let project env sorts posn t1 t2 = let t1 = EConstr.Unsafe.to_constr t1 in let t2 = EConstr.Unsafe.to_constr t2 in - let ty1 = get_type_of env sigma t1 in - let s = get_sort_family_of env sigma ty1 in + let ty1 = get_type_of env sigma (EConstr.of_constr t1) in + let s = get_sort_family_of env sigma (EConstr.of_constr ty1) in if Sorts.List.mem s sorts then [(List.rev posn,t1,t2)] else [] in @@ -842,7 +842,7 @@ let injectable env sigma t1 t2 = let descend_then env sigma head dirn = let IndType (indf,_) = - try find_rectype env sigma (EConstr.of_constr (get_type_of env sigma head)) + try find_rectype env sigma (EConstr.of_constr (get_type_of env sigma (EConstr.of_constr head))) with Not_found -> error "Cannot project on an inductive type derived from a dependency." in let indp,_ = (dest_ind_family indf) in @@ -897,7 +897,7 @@ let build_selector env sigma dirn c ind special default = dependent types.") in let (indp,_) = dest_ind_family indf in let ind, _ = check_privacy env indp in - let typ = Retyping.get_type_of env sigma default in + let typ = Retyping.get_type_of env sigma (EConstr.of_constr default) in let (mib,mip) = lookup_mind_specif env ind in let deparsign = make_arity_signature env true indf in let p = it_mkLambda_or_LetIn typ deparsign in @@ -912,7 +912,7 @@ let build_selector env sigma dirn c ind special default = let rec build_discriminator env sigma dirn c = function | [] -> - let ind = get_type_of env sigma c in + let ind = get_type_of env sigma (EConstr.of_constr c) in let true_0,false_0 = build_coq_True(),build_coq_False() in build_selector env sigma dirn c ind true_0 false_0 @@ -1084,7 +1084,7 @@ let find_sigma_data env s = build_sigma_type () let make_tuple env sigma (rterm,rty) lind = assert (not (EConstr.Vars.noccurn sigma lind (EConstr.of_constr rty))); - let sigdata = find_sigma_data env (get_sort_of env sigma rty) in + let sigdata = find_sigma_data env (get_sort_of env sigma (EConstr.of_constr rty)) in let sigma, a = type_of ~refresh:true env sigma (mkRel lind) in let na = Context.Rel.Declaration.get_name (lookup_rel lind env) in (* We move [lind] to [1] and lift other rels > [lind] by 1 *) @@ -1262,7 +1262,7 @@ let sig_clausal_form env sigma sort_of_ty siglen ty dflt = let make_iterated_tuple env sigma dflt (z,zty) = let (zty,rels) = minimal_free_rels_rec env sigma (z,zty) in - let sort_of_zty = get_sort_of env sigma zty in + let sort_of_zty = get_sort_of env sigma (EConstr.of_constr zty) in let sorted_rels = Int.Set.elements rels in let sigma, (tuple,tuplety) = List.fold_left (fun (sigma, t) -> make_tuple env sigma t) (sigma, (z,zty)) sorted_rels @@ -1533,7 +1533,7 @@ let decomp_tuple_term env sigma c t = let subst_tuple_term env sigma dep_pair1 dep_pair2 b = let sigma = Sigma.to_evar_map sigma in - let typ = get_type_of env sigma dep_pair1 in + let typ = get_type_of env sigma (EConstr.of_constr dep_pair1) in (* We find all possible decompositions *) let decomps1 = decomp_tuple_term env sigma dep_pair1 typ in let decomps2 = decomp_tuple_term env sigma dep_pair2 typ in @@ -1623,7 +1623,7 @@ let cutRewriteInConcl l2r eqn = cutRewriteClause l2r eqn None let substClause l2r c cls = Proofview.Goal.enter { enter = begin fun gl -> - let eq = pf_apply get_type_of gl c in + let eq = pf_apply get_type_of gl (EConstr.of_constr c) in tclTHENS (cutSubstClause l2r eq cls) [Proofview.tclUNIT (); exact_no_check c] end } diff --git a/tactics/hints.ml b/tactics/hints.ml index c41f88ab7..2aa434777 100644 --- a/tactics/hints.ml +++ b/tactics/hints.ml @@ -846,7 +846,7 @@ let fresh_global_or_constr env sigma poly cr = let make_resolves env sigma flags pri poly ?name cr = let c, ctx = fresh_global_or_constr env sigma poly cr in - let cty = Retyping.get_type_of env sigma c in + let cty = Retyping.get_type_of env sigma (EConstr.of_constr c) in let try_apply f = try Some (f (c, cty, ctx)) with Failure _ -> None in let ents = List.map_filter try_apply diff --git a/tactics/inv.ml b/tactics/inv.ml index 0b2d2f0b2..38f75995b 100644 --- a/tactics/inv.ml +++ b/tactics/inv.ml @@ -65,7 +65,7 @@ let var_occurs_in_pf gl id = type inversion_status = Dep of constr option | NoDep let compute_eqn env sigma n i ai = - (mkRel (n-i),get_type_of env sigma (mkRel (n-i))) + (mkRel (n-i),get_type_of env sigma (EConstr.of_constr (mkRel (n-i)))) let make_inv_predicate env evd indf realargs id status concl = let nrealargs = List.length realargs in @@ -86,7 +86,7 @@ let make_inv_predicate env evd indf realargs id status concl = match dflt_concl with | Some concl -> concl (*assumed it's some [x1..xn,H:I(x1..xn)]C*) | None -> - let sort = get_sort_family_of env !evd concl in + let sort = get_sort_family_of env !evd (EConstr.of_constr concl) in let sort = Evarutil.evd_comb1 (Evd.fresh_sort_in_family env) evd sort in let p = make_arity env true indf sort in let evd',(p,ptyp) = Unification.abstract_list_all env diff --git a/tactics/tacticals.ml b/tactics/tacticals.ml index 676b23d09..2754db010 100644 --- a/tactics/tacticals.ml +++ b/tactics/tacticals.ml @@ -248,10 +248,10 @@ let compute_constructor_signatures isrec ((_,k as ity),u) = Array.map2 analrec lc lrecargs let elimination_sort_of_goal gl = - pf_apply Retyping.get_sort_family_of gl (pf_concl gl) + pf_apply Retyping.get_sort_family_of gl (EConstr.of_constr (pf_concl gl)) let elimination_sort_of_hyp id gl = - pf_apply Retyping.get_sort_family_of gl (pf_get_hyp_typ gl id) + pf_apply Retyping.get_sort_family_of gl (EConstr.of_constr (pf_get_hyp_typ gl id)) let elimination_sort_of_clause = function | None -> elimination_sort_of_goal @@ -708,12 +708,12 @@ module New = struct let elimination_sort_of_goal gl = (** Retyping will expand evars anyway. *) let c = Proofview.Goal.concl (Goal.assume gl) in - pf_apply Retyping.get_sort_family_of gl c + pf_apply Retyping.get_sort_family_of gl (EConstr.of_constr c) let elimination_sort_of_hyp id gl = (** Retyping will expand evars anyway. *) let c = pf_get_hyp_typ id (Goal.assume gl) in - pf_apply Retyping.get_sort_family_of gl c + pf_apply Retyping.get_sort_family_of gl (EConstr.of_constr c) let elimination_sort_of_clause id gl = match id with | None -> elimination_sort_of_goal gl diff --git a/tactics/tactics.ml b/tactics/tactics.ml index c96553fae..e294f928e 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -412,7 +412,7 @@ let default_id env sigma decl = let open Context.Rel.Declaration in match decl with | LocalAssum (name,t) -> - let dft = default_id_of_sort (Retyping.get_sort_of env sigma t) in + let dft = default_id_of_sort (Retyping.get_sort_of env sigma (EConstr.of_constr t)) in id_of_name_with_default dft name | LocalDef (name,b,_) -> id_of_name_using_hdchar env b name @@ -784,9 +784,9 @@ let make_change_arg c pats = { run = fun sigma -> Sigma.here (replace_vars (Id.Map.bindings pats) c) sigma } let check_types env sigma mayneedglobalcheck deep newc origc = - let t1 = Retyping.get_type_of env sigma newc in + let t1 = Retyping.get_type_of env sigma (EConstr.of_constr newc) in if deep then begin - let t2 = Retyping.get_type_of env sigma origc in + let t2 = Retyping.get_type_of env sigma (EConstr.of_constr origc) in let sigma, t2 = Evarsolve.refresh_universes ~onlyalg:true (Some false) env sigma t2 in let sigma, b = infer_conv ~pb:Reduction.CUMUL env sigma t1 t2 in @@ -1341,7 +1341,7 @@ let enforce_prop_bound_names rename tac = | Prod (Name _ as na,t,t') -> let very_standard = true in let na = - if Retyping.get_sort_family_of env sigma t = InProp then + if Retyping.get_sort_family_of env sigma (EConstr.of_constr t) = InProp then (* "very_standard" says that we should have "H" names only, but this would break compatibility even more... *) let s = match Namegen.head_name t with @@ -1411,7 +1411,7 @@ let general_elim_clause_gen elimtac indclause elim = let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in let (elimc,lbindelimc) = elim.elimbody in - let elimt = Retyping.get_type_of env sigma elimc in + let elimt = Retyping.get_type_of env sigma (EConstr.of_constr elimc) in let i = match elim.elimindex with None -> index_of_ind_arg elimt | Some i -> i in elimtac elim.elimrename i (elimc, elimt, lbindelimc) indclause @@ -1421,7 +1421,7 @@ let general_elim with_evars clear_flag (c, lbindc) elim = Proofview.Goal.enter { enter = begin fun gl -> let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in - let ct = Retyping.get_type_of env sigma c in + let ct = Retyping.get_type_of env sigma (EConstr.of_constr c) in let t = try snd (reduce_to_quantified_ind env sigma ct) with UserError _ -> ct in let elimtac = elimination_clause_scheme with_evars in let indclause = make_clenv_binding env sigma (c, t) lbindc in @@ -1439,7 +1439,7 @@ let general_case_analysis_in_context with_evars clear_flag (c,lbindc) = let sigma = Proofview.Goal.sigma gl in let env = Proofview.Goal.env gl in let concl = Proofview.Goal.concl gl in - let t = Retyping.get_type_of env (Sigma.to_evar_map sigma) c in + let t = Retyping.get_type_of env (Sigma.to_evar_map sigma) (EConstr.of_constr c) in let (mind,_) = reduce_to_quantified_ind env (Sigma.to_evar_map sigma) t in let sort = Tacticals.New.elimination_sort_of_goal gl in let Sigma (elim, sigma, p) = @@ -1554,7 +1554,7 @@ let elimination_in_clause_scheme with_evars ?(flags=elim_flags ()) (str "The type of elimination clause is not well-formed.") in let elimclause' = clenv_fchain ~flags indmv elimclause indclause in let hyp = mkVar id in - let hyp_typ = Retyping.get_type_of env sigma hyp in + let hyp_typ = Retyping.get_type_of env sigma (EConstr.of_constr hyp) in let hypclause = mk_clenv_from_env env sigma (Some 0) (hyp, hyp_typ) in let elimclause'' = clenv_fchain_in id ~flags hypmv elimclause' hypclause in let new_hyp_typ = clenv_type elimclause'' in @@ -1614,7 +1614,7 @@ let make_projection env sigma params cstr sign elim i n c u = [|mkApp (c, args)|]) in let app = it_mkLambda_or_LetIn proj sign in - let t = Retyping.get_type_of env sigma app in + let t = Retyping.get_type_of env sigma (EConstr.of_constr app) in Some (app, t) | None -> None in elim @@ -1624,7 +1624,7 @@ let descend_in_conjunctions avoid tac (err, info) c = let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in try - let t = Retyping.get_type_of env sigma c in + let t = Retyping.get_type_of env sigma (EConstr.of_constr c) in let ((ind,u),t) = reduce_to_quantified_ind env sigma t in let sign,ccl = decompose_prod_assum t in match match_with_tuple sigma ccl with @@ -1704,7 +1704,7 @@ let general_apply with_delta with_destruct with_evars clear_flag (loc,(c,lbind)) let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in - let thm_ty0 = nf_betaiota sigma (EConstr.of_constr (Retyping.get_type_of env sigma c)) in + let thm_ty0 = nf_betaiota sigma (EConstr.of_constr (Retyping.get_type_of env sigma (EConstr.of_constr c))) in let try_apply thm_ty nprod = try let n = nb_prod_modulo_zeta sigma (EConstr.of_constr thm_ty) - nprod in @@ -1830,7 +1830,7 @@ let progress_with_clause flags innerclause clause = with Not_found -> error "Unable to unify." let apply_in_once_main flags innerclause env sigma (d,lbind) = - let thm = nf_betaiota sigma (EConstr.of_constr (Retyping.get_type_of env sigma d)) in + let thm = nf_betaiota sigma (EConstr.of_constr (Retyping.get_type_of env sigma (EConstr.of_constr d))) in let rec aux clause = try progress_with_clause flags innerclause clause with e when CErrors.noncritical e -> @@ -2604,7 +2604,7 @@ let letin_tac_gen with_eq (id,depdecls,lastlhyp,ccl,c) ty = let Sigma (t, sigma, p) = match ty with | Some t -> Sigma.here t sigma | None -> - let t = typ_of env sigma c in + let t = typ_of env sigma (EConstr.of_constr c) in let sigma, c = Evarsolve.refresh_universes ~onlyalg:true (Some false) env (Sigma.to_evar_map sigma) t in Sigma.Unsafe.of_pair (c, sigma) in @@ -2656,7 +2656,7 @@ let insert_before decls lasthyp env = let mkletin_goal env sigma store with_eq dep (id,lastlhyp,ccl,c) ty = let open Context.Named.Declaration in - let t = match ty with Some t -> t | _ -> typ_of env sigma c in + let t = match ty with Some t -> t | _ -> typ_of env sigma (EConstr.of_constr c) in let decl = if dep then LocalDef (id,c,t) else LocalAssum (id,t) in @@ -2903,7 +2903,7 @@ let specialize (c,lbind) ipat = let sigma = Typeclasses.resolve_typeclasses env sigma in sigma, nf_evar sigma c else - let clause = make_clenv_binding env sigma (c,Retyping.get_type_of env sigma c) lbind in + let clause = make_clenv_binding env sigma (c,Retyping.get_type_of env sigma (EConstr.of_constr c)) lbind in let flags = { (default_unify_flags ()) with resolve_evars = true } in let clause = clenv_unify_meta_types ~flags clause in let (thd,tstack) = whd_nored_stack clause.evd (EConstr.of_constr (clenv_value clause)) in @@ -2919,7 +2919,7 @@ let specialize (c,lbind) ipat = pr_name (meta_name clause.evd (List.hd (collect_metas clause.evd (EConstr.of_constr term)))) ++ str "."); clause.evd, term in - let typ = Retyping.get_type_of env sigma term in + let typ = Retyping.get_type_of env sigma (EConstr.of_constr term) in let tac = match kind_of_term (fst(decompose_app (snd(decompose_lam_assum c)))) with | Var id when Id.List.mem id (Tacmach.New.pf_ids_of_hyps gl) -> @@ -3152,7 +3152,7 @@ let expand_projections env sigma c = let sigma = Sigma.to_evar_map sigma in let rec aux env c = match EConstr.kind sigma c with - | Proj (p, c) -> EConstr.of_constr (Retyping.expand_projection env sigma p (EConstr.Unsafe.to_constr (aux env c)) []) + | Proj (p, c) -> EConstr.of_constr (Retyping.expand_projection env sigma p (aux env c) []) | _ -> map_constr_with_full_binders sigma push_rel aux env c in EConstr.Unsafe.to_constr (aux env (EConstr.of_constr c)) @@ -3673,7 +3673,7 @@ let abstract_args gl generalize_vars dep id defined f args = else [] in let body, c' = - if defined then Some c', Retyping.get_type_of ctxenv !sigma c' + if defined then Some c', Retyping.get_type_of ctxenv !sigma (EConstr.of_constr c') else None, c' in let typ = Tacmach.pf_get_hyp_typ gl id in @@ -4132,7 +4132,7 @@ let apply_induction_in_context with_evars hyp0 inhyps elim indvars names induct_ let dep_in_concl = Option.cata (fun id -> occur_var env sigma id (EConstr.of_constr concl)) false hyp0 in let dep = dep_in_hyps || dep_in_concl in let tmpcl = it_mkNamedProd_or_LetIn concl deps in - let s = Retyping.get_sort_family_of env sigma tmpcl in + let s = Retyping.get_sort_family_of env sigma (EConstr.of_constr tmpcl) in let deps_cstr = List.fold_left (fun a decl -> if NamedDecl.is_local_assum decl then (mkVar (NamedDecl.get_id decl))::a else a) [] deps in @@ -4286,7 +4286,7 @@ let check_enough_applied env sigma elim = fun u -> let t,_ = decompose_app (whd_all env sigma (EConstr.of_constr u)) in isInd t | Some elimc -> - let elimt = Retyping.get_type_of env sigma (fst elimc) in + let elimt = Retyping.get_type_of env sigma (EConstr.of_constr (fst elimc)) in let scheme = compute_elim_sig ~elimc elimt in match scheme.indref with | None -> @@ -4331,7 +4331,7 @@ let pose_induction_arg_then isrec with_evars (is_arg_pure_hyp,from_prefix) elim Refine.refine ~unsafe:true { run = begin fun sigma -> let b = not with_evars && with_eq != None in let Sigma (c, sigma, p) = use_bindings env sigma elim b (c0,lbind) t0 in - let t = Retyping.get_type_of env (Sigma.to_evar_map sigma) c in + let t = Retyping.get_type_of env (Sigma.to_evar_map sigma) (EConstr.of_constr c) in let Sigma (ans, sigma, q) = mkletin_goal env sigma store with_eq false (id,lastlhyp,ccl,c) (Some t) in Sigma (ans, sigma, p +> q) end }; @@ -4376,7 +4376,7 @@ let induction_gen clear_flag isrec with_evars elim let sigma = Proofview.Goal.sigma gl in let ccl = Proofview.Goal.raw_concl gl in let cls = Option.default allHypsAndConcl cls in - let t = typ_of env sigma c in + let t = typ_of env sigma (EConstr.of_constr c) in let is_arg_pure_hyp = isVar c && not (mem_named_context_val (destVar c) (Global.named_context_val ())) && lbind == NoBindings && not with_evars && Option.is_empty eqname diff --git a/toplevel/command.ml b/toplevel/command.ml index 249d41845..14a45f837 100644 --- a/toplevel/command.ml +++ b/toplevel/command.ml @@ -206,7 +206,7 @@ let do_definition ident k pl bl red_option c ctypopt hook = assert(Univ.ContextSet.is_empty ctx); let typ = match ce.const_entry_type with | Some t -> t - | None -> Retyping.get_type_of env evd c + | None -> Retyping.get_type_of env evd (EConstr.of_constr c) in Obligations.check_evars env evd; let obls, _, c, cty = @@ -459,7 +459,7 @@ let sign_level env evd sign = | LocalDef _ -> lev, push_rel d env | LocalAssum _ -> let s = destSort (Reduction.whd_all env - (nf_evar evd (Retyping.get_type_of env evd (RelDecl.get_type d)))) + (nf_evar evd (Retyping.get_type_of env evd (EConstr.of_constr (RelDecl.get_type d))))) in let u = univ_of_sort s in (Univ.sup u lev, push_rel d env)) diff --git a/toplevel/indschemes.ml b/toplevel/indschemes.ml index 3f818f960..fef99d8af 100644 --- a/toplevel/indschemes.ml +++ b/toplevel/indschemes.ml @@ -403,7 +403,7 @@ let do_mutual_induction_scheme lnamedepindsort = in let sigma, listdecl = Indrec.build_mutual_induction_scheme env0 sigma lrecspec in let declare decl fi lrecref = - let decltype = Retyping.get_type_of env0 sigma decl in + let decltype = Retyping.get_type_of env0 sigma (EConstr.of_constr decl) in let proof_output = Future.from_val ((decl,Univ.ContextSet.empty),Safe_typing.empty_private_constants) in let cst = define fi UserIndividualRequest sigma proof_output (Some decltype) in ConstRef cst :: lrecref diff --git a/toplevel/record.ml b/toplevel/record.ml index ffe7980ef..5f2b99f90 100644 --- a/toplevel/record.ml +++ b/toplevel/record.ml @@ -84,7 +84,7 @@ let compute_constructor_level evars env l = List.fold_right (fun d (env, univ) -> let univ = if is_local_assum d then - let s = Retyping.get_sort_of env evars (RelDecl.get_type d) in + let s = Retyping.get_sort_of env evars (EConstr.of_constr (RelDecl.get_type d)) in Univ.sup (univ_of_sort s) univ else univ in (push_rel d env, univ)) diff --git a/toplevel/vernacentries.ml b/toplevel/vernacentries.ml index 593aa9578..c39b4dc25 100644 --- a/toplevel/vernacentries.ml +++ b/toplevel/vernacentries.ml @@ -1579,7 +1579,7 @@ let vernac_check_may_eval redexp glopt rc = let c = nf c in let j = if Evarutil.has_undefined_evars sigma' c then - Evarutil.j_nf_evar sigma' (Retyping.get_judgment_of env sigma' c) + Evarutil.j_nf_evar sigma' (Retyping.get_judgment_of env sigma' (EConstr.of_constr c)) else (* OK to call kernel which does not support evars *) Arguments_renaming.rename_typing env c in -- cgit v1.2.3 From 214a2ffd2cce3fa24908710af7db528a40345db6 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sat, 5 Nov 2016 18:24:20 +0100 Subject: Cbv API using EConstr. --- pretyping/cbv.ml | 1 + pretyping/cbv.mli | 2 +- pretyping/tacred.ml | 2 +- 3 files changed, 3 insertions(+), 2 deletions(-) diff --git a/pretyping/cbv.ml b/pretyping/cbv.ml index 84bf849e7..a42061f28 100644 --- a/pretyping/cbv.ml +++ b/pretyping/cbv.ml @@ -376,6 +376,7 @@ and cbv_norm_value info = function (* reduction under binders *) (* with profiling *) let cbv_norm infos constr = + let constr = EConstr.Unsafe.to_constr constr in with_stats (lazy (cbv_norm_term infos (subs_id 0) constr)) type cbv_infos = cbv_value infos diff --git a/pretyping/cbv.mli b/pretyping/cbv.mli index 87a03abbd..3d1745767 100644 --- a/pretyping/cbv.mli +++ b/pretyping/cbv.mli @@ -19,7 +19,7 @@ open Esubst type cbv_infos val create_cbv_infos : RedFlags.reds -> env -> Evd.evar_map -> cbv_infos -val cbv_norm : cbv_infos -> constr -> constr +val cbv_norm : cbv_infos -> EConstr.constr -> constr (*********************************************************************** i This is for cbv debug *) diff --git a/pretyping/tacred.ml b/pretyping/tacred.ml index ac2a3bc49..c1e9573d2 100644 --- a/pretyping/tacred.ml +++ b/pretyping/tacred.ml @@ -1136,7 +1136,7 @@ let fold_commands cl env sigma c = (* call by value reduction functions *) let cbv_norm_flags flags env sigma t = - cbv_norm (create_cbv_infos flags env sigma) (EConstr.Unsafe.to_constr t) + cbv_norm (create_cbv_infos flags env sigma) t let cbv_beta = cbv_norm_flags beta empty_env let cbv_betaiota = cbv_norm_flags betaiota empty_env -- cgit v1.2.3 From 83607f75a13ea915affa8cfc5bfc14cc944c61ef Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sat, 5 Nov 2016 18:45:55 +0100 Subject: Find_subterm API using EConstr. --- pretyping/find_subterm.ml | 25 +++++++++++++++---------- pretyping/find_subterm.mli | 12 ++++++------ pretyping/pretype_errors.ml | 2 +- pretyping/pretype_errors.mli | 2 +- pretyping/tacred.ml | 7 ++++--- pretyping/unification.ml | 23 ++++++++++++----------- tactics/tactics.ml | 2 +- toplevel/himsg.ml | 4 ++-- 8 files changed, 42 insertions(+), 35 deletions(-) diff --git a/pretyping/find_subterm.ml b/pretyping/find_subterm.ml index 4b9cf415f..d7f2d54aa 100644 --- a/pretyping/find_subterm.ml +++ b/pretyping/find_subterm.ml @@ -11,7 +11,7 @@ open Util open CErrors open Names open Locus -open Term +open EConstr open Nameops open Termops open Pretype_errors @@ -63,7 +63,10 @@ let proceed_with_occurrences f occs x = let map_named_declaration_with_hyploc f hyploc acc decl = let open Context.Named.Declaration in - let f = f (Some (NamedDecl.get_id decl, hyploc)) in + let f acc typ = + let acc, typ = f (Some (NamedDecl.get_id decl, hyploc)) acc (EConstr.of_constr typ) in + acc, EConstr.Unsafe.to_constr typ + in match decl,hyploc with | LocalAssum (id,_), InHypValueOnly -> error_occurrences_error (IncorrectInValueOccurrence id) @@ -82,10 +85,10 @@ let map_named_declaration_with_hyploc f hyploc acc decl = exception SubtermUnificationError of subterm_unification_error -exception NotUnifiable of (constr * constr * unification_error) option +exception NotUnifiable of (Constr.t * Constr.t * unification_error) option type 'a testing_function = { - match_fun : 'a -> constr -> 'a; + match_fun : 'a -> EConstr.constr -> 'a; merge_fun : 'a -> 'a -> 'a; mutable testing_state : 'a; mutable last_found : position_reporting option @@ -130,7 +133,8 @@ let replace_term_occ_gen_modulo occs like_first test bywhat cl occ t = with NotUnifiable _ -> subst_below k t and subst_below k t = - map_constr_with_binders_left_to_right (fun d k -> k+1) substrec k t + let substrec i c = EConstr.Unsafe.to_constr (substrec i (EConstr.of_constr c)) in + EConstr.of_constr (map_constr_with_binders_left_to_right (fun d k -> k+1) substrec k (EConstr.Unsafe.to_constr t)) in let t' = substrec 0 t in (!pos, t') @@ -138,8 +142,8 @@ let replace_term_occ_gen_modulo occs like_first test bywhat cl occ t = let replace_term_occ_modulo occs test bywhat t = let occs',like_first = match occs with AtOccs occs -> occs,false | LikeFirst -> AllOccurrences,true in - proceed_with_occurrences - (replace_term_occ_gen_modulo occs' like_first test bywhat None) occs' t + EConstr.Unsafe.to_constr (proceed_with_occurrences + (replace_term_occ_gen_modulo occs' like_first test bywhat None) occs' t) let replace_term_occ_decl_modulo occs test bywhat d = let (plocs,hyploc),like_first = @@ -154,11 +158,12 @@ let replace_term_occ_decl_modulo occs test bywhat d = let make_eq_univs_test env evd c = { match_fun = (fun evd c' -> - let b, cst = Universes.eq_constr_universes_proj env c c' in - if b then + match EConstr.eq_constr_universes_proj env evd c c' with + | None -> raise (NotUnifiable None) + | Some cst -> try Evd.add_universe_constraints evd cst with Evd.UniversesDiffer -> raise (NotUnifiable None) - else raise (NotUnifiable None)); + ); merge_fun = (fun evd _ -> evd); testing_state = evd; last_found = None diff --git a/pretyping/find_subterm.mli b/pretyping/find_subterm.mli index c741ab048..49a5dd7f2 100644 --- a/pretyping/find_subterm.mli +++ b/pretyping/find_subterm.mli @@ -26,7 +26,7 @@ exception SubtermUnificationError of subterm_unification_error with None. *) type 'a testing_function = { - match_fun : 'a -> constr -> 'a; + match_fun : 'a -> EConstr.constr -> 'a; merge_fun : 'a -> 'a -> 'a; mutable testing_state : 'a; mutable last_found : position_reporting option @@ -34,7 +34,7 @@ type 'a testing_function = { (** This is the basic testing function, looking for exact matches of a closed term *) -val make_eq_univs_test : env -> evar_map -> constr -> evar_map testing_function +val make_eq_univs_test : env -> evar_map -> EConstr.constr -> evar_map testing_function (** [replace_term_occ_modulo occl test mk c] looks in [c] for subterm modulo a testing function [test] and replaces successfully @@ -42,26 +42,26 @@ val make_eq_univs_test : env -> evar_map -> constr -> evar_map testing_function ()]; it turns a NotUnifiable exception raised by the testing function into a SubtermUnificationError. *) val replace_term_occ_modulo : occurrences or_like_first -> - 'a testing_function -> (unit -> constr) -> constr -> constr + 'a testing_function -> (unit -> EConstr.constr) -> EConstr.constr -> constr (** [replace_term_occ_decl_modulo] is similar to [replace_term_occ_modulo] but for a named_declaration. *) val replace_term_occ_decl_modulo : (occurrences * hyp_location_flag) or_like_first -> - 'a testing_function -> (unit -> constr) -> + 'a testing_function -> (unit -> EConstr.constr) -> Context.Named.Declaration.t -> Context.Named.Declaration.t (** [subst_closed_term_occ occl c d] replaces occurrences of closed [c] at positions [occl] by [Rel 1] in [d] (see also Note OCC), unifying universes which results in a set of constraints. *) val subst_closed_term_occ : env -> evar_map -> occurrences or_like_first -> - constr -> constr -> constr * evar_map + EConstr.constr -> EConstr.constr -> constr * evar_map (** [subst_closed_term_occ_decl evd occl c decl] replaces occurrences of closed [c] at positions [occl] by [Rel 1] in [decl]. *) val subst_closed_term_occ_decl : env -> evar_map -> (occurrences * hyp_location_flag) or_like_first -> - constr -> Context.Named.Declaration.t -> Context.Named.Declaration.t * evar_map + EConstr.constr -> Context.Named.Declaration.t -> Context.Named.Declaration.t * evar_map (** Miscellaneous *) val error_invalid_occurrence : int list -> 'a diff --git a/pretyping/pretype_errors.ml b/pretyping/pretype_errors.ml index 5b0958695..f8f6d44bf 100644 --- a/pretyping/pretype_errors.ml +++ b/pretyping/pretype_errors.ml @@ -27,7 +27,7 @@ type unification_error = type position = (Id.t * Locus.hyp_location_flag) option -type position_reporting = (position * int) * constr +type position_reporting = (position * int) * EConstr.t type subterm_unification_error = bool * position_reporting * position_reporting * (constr * constr * unification_error) option diff --git a/pretyping/pretype_errors.mli b/pretyping/pretype_errors.mli index 73f81923f..b015add79 100644 --- a/pretyping/pretype_errors.mli +++ b/pretyping/pretype_errors.mli @@ -28,7 +28,7 @@ type unification_error = type position = (Id.t * Locus.hyp_location_flag) option -type position_reporting = (position * int) * constr +type position_reporting = (position * int) * EConstr.t type subterm_unification_error = bool * position_reporting * position_reporting * (constr * constr * unification_error) option diff --git a/pretyping/tacred.ml b/pretyping/tacred.ml index c1e9573d2..290d77b1b 100644 --- a/pretyping/tacred.ml +++ b/pretyping/tacred.ml @@ -1150,13 +1150,14 @@ let compute = cbv_betadeltaiota * the specified occurrences. *) let abstract_scheme env (locc,a) (c, sigma) = - let ta = Retyping.get_type_of env sigma (EConstr.of_constr a) in + let a = EConstr.of_constr a in + let ta = Retyping.get_type_of env sigma a in let na = named_hd env ta Anonymous in if occur_meta sigma (EConstr.of_constr ta) then error "Cannot find a type for the generalisation."; - if occur_meta sigma (EConstr.of_constr a) then + if occur_meta sigma a then mkLambda (na,ta,c), sigma else - let c', sigma' = subst_closed_term_occ env sigma (AtOccs locc) a c in + let c', sigma' = subst_closed_term_occ env sigma (AtOccs locc) a (EConstr.of_constr c) in mkLambda (na,ta,c'), sigma' let pattern_occs loccs_trm = { e_redfun = begin fun env sigma c -> diff --git a/pretyping/unification.ml b/pretyping/unification.ml index 848a2f4a8..8865e69ef 100644 --- a/pretyping/unification.ml +++ b/pretyping/unification.ml @@ -89,7 +89,7 @@ let abstract_scheme env evd c l lname_typ = else *) if occur_meta evd (EConstr.of_constr a) then mkLambda_name env (na,ta,t), evd else - let t', evd' = Find_subterm.subst_closed_term_occ env evd locc a t in + let t', evd' = Find_subterm.subst_closed_term_occ env evd locc (EConstr.of_constr a) (EConstr.of_constr t) in mkLambda_name env (na,ta,t'), evd') (c,evd) (List.rev l) @@ -1544,20 +1544,21 @@ let make_pattern_test from_prefix_of_ind is_correct_type env sigma (pending,c) = else default_matching_flags pending in let n = List.length (snd (decompose_app c)) in let matching_fun _ t = + let open EConstr in try let t',l2 = if from_prefix_of_ind then (* We check for fully applied subterms of the form "u u1 .. un" *) (* of inductive type knowing only a prefix "u u1 .. ui" *) - let t,l = decompose_app t in + let t,l = decompose_app sigma t in let l1,l2 = try List.chop n l with Failure _ -> raise (NotUnifiable None) in - if not (List.for_all closed0 l2) then raise (NotUnifiable None) + if not (List.for_all (fun c -> Vars.closed0 sigma c) l2) then raise (NotUnifiable None) else applist (t,l1), l2 else t, [] in - let sigma = w_typed_unify env sigma Reduction.CONV flags c t' in - let ty = Retyping.get_type_of env sigma (EConstr.of_constr t) in + let sigma = w_typed_unify env sigma Reduction.CONV flags c (EConstr.Unsafe.to_constr t') in + let ty = Retyping.get_type_of env sigma t in if not (is_correct_type ty) then raise (NotUnifiable None); Some(sigma, t, l2) with @@ -1568,7 +1569,7 @@ let make_pattern_test from_prefix_of_ind is_correct_type env sigma (pending,c) = let merge_fun c1 c2 = match c1, c2 with | Some (evd,c1,x), Some (_,c2,_) -> - let (evd,b) = infer_conv ~pb:CONV env evd c1 c2 in + let (evd,b) = infer_conv ~pb:CONV env evd (EConstr.Unsafe.to_constr c1) (EConstr.Unsafe.to_constr c2) in if b then Some (evd, c1, x) else raise (NotUnifiable None) | Some _, None -> c1 | None, Some _ -> c2 @@ -1578,13 +1579,13 @@ let make_pattern_test from_prefix_of_ind is_correct_type env sigma (pending,c) = (fun test -> match test.testing_state with | None -> None | Some (sigma,_,l) -> - let c = applist (nf_evar sigma (local_strong whd_meta sigma (EConstr.of_constr c)),l) in + let c = applist (nf_evar sigma (local_strong whd_meta sigma (EConstr.of_constr c)), List.map (EConstr.to_constr sigma) l) in let univs, subst = nf_univ_variables sigma in Some (sigma,subst_univs_constr subst c)) let make_eq_test env evd c = let out cstr = - match cstr.last_found with None -> None | _ -> Some (cstr.testing_state, c) + match cstr.last_found with None -> None | _ -> Some (cstr.testing_state, EConstr.Unsafe.to_constr c) in (make_eq_univs_test env evd c, out) @@ -1601,7 +1602,7 @@ let make_abstraction_core name (test,out) env sigma c ty occs check_occs concl = x in let likefirst = clause_with_generic_occurrences occs in - let mkvarid () = mkVar id in + let mkvarid () = EConstr.mkVar id in let compute_dependency _ d (sign,depdecls) = let hyp = NamedDecl.get_id d in match occurrences_of_hyp hyp occs with @@ -1630,7 +1631,7 @@ let make_abstraction_core name (test,out) env sigma c ty occs check_occs concl = | NoOccurrences -> concl | occ -> let occ = if likefirst then LikeFirst else AtOccs occ in - replace_term_occ_modulo occ test mkvarid concl + replace_term_occ_modulo occ test mkvarid (EConstr.of_constr concl) in let lastlhyp = if List.is_empty depdecls then None else Some (NamedDecl.get_id (List.last depdecls)) in @@ -1674,7 +1675,7 @@ let make_abstraction env evd ccl abs = env evd (snd c) None occs check_occs ccl | AbstractExact (name,c,ty,occs,check_occs) -> make_abstraction_core name - (make_eq_test env evd c) + (make_eq_test env evd (EConstr.of_constr c)) env evd c ty occs check_occs ccl let keyed_unify env evd kop = diff --git a/tactics/tactics.ml b/tactics/tactics.ml index e294f928e..22d01e401 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -2772,7 +2772,7 @@ let generalize_goal_gen env sigma ids i ((occs,c,b),na) t cl = let decls,cl = decompose_prod_n_assum i cl in let dummy_prod = EConstr.of_constr (it_mkProd_or_LetIn mkProp decls) in let newdecls,_ = decompose_prod_n_assum i (subst_term_gen sigma EConstr.eq_constr_nounivs (EConstr.of_constr c) dummy_prod) in - let cl',sigma' = subst_closed_term_occ env sigma (AtOccs occs) c (it_mkProd_or_LetIn cl newdecls) in + let cl',sigma' = subst_closed_term_occ env sigma (AtOccs occs) (EConstr.of_constr c) (EConstr.of_constr (it_mkProd_or_LetIn cl newdecls)) in let na = generalized_name c t ids cl' na in let decl = match b with | None -> LocalAssum (na,t) diff --git a/toplevel/himsg.ml b/toplevel/himsg.ml index 0279ff0c5..4cffbee82 100644 --- a/toplevel/himsg.ml +++ b/toplevel/himsg.ml @@ -734,10 +734,10 @@ let explain_cannot_unify_occurrences env sigma nested ((cl2,pos2),t2) ((cl1,pos1 else let ppreason = match e with None -> mt() | Some (c1,c2,e) -> explain_unification_error env sigma c1 c2 (Some e) in str "Found incompatible occurrences of the pattern" ++ str ":" ++ - spc () ++ str "Matched term " ++ pr_lconstr_env env sigma t2 ++ + spc () ++ str "Matched term " ++ pr_lconstr_env env sigma (EConstr.to_constr sigma t2) ++ strbrk " at position " ++ pr_position (cl2,pos2) ++ strbrk " is not compatible with matched term " ++ - pr_lconstr_env env sigma t1 ++ strbrk " at position " ++ + pr_lconstr_env env sigma (EConstr.to_constr sigma t1) ++ strbrk " at position " ++ pr_position (cl1,pos1) ++ ppreason ++ str "." let pr_constraints printenv env sigma evars cstrs = -- cgit v1.2.3 From 147afe827dc83602cc160a8b1357e84ecea910ff Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sat, 5 Nov 2016 20:13:32 +0100 Subject: Evardefine API using EConstr. --- pretyping/cases.ml | 8 +++--- pretyping/coercion.ml | 11 ++++---- pretyping/evarconv.ml | 2 +- pretyping/evardefine.ml | 65 ++++++++++++++++++++++++-------------------- pretyping/evardefine.mli | 12 ++++---- pretyping/pretype_errors.ml | 4 +-- pretyping/pretype_errors.mli | 8 +++--- pretyping/pretyping.ml | 20 +++++++------- pretyping/typing.ml | 21 +++++++------- pretyping/unification.ml | 11 ++++---- toplevel/himsg.ml | 4 +-- 11 files changed, 88 insertions(+), 78 deletions(-) diff --git a/pretyping/cases.ml b/pretyping/cases.ml index 4dd502106..915cd261d 100644 --- a/pretyping/cases.ml +++ b/pretyping/cases.ml @@ -336,7 +336,7 @@ let unify_tomatch_with_patterns evdref env loc typ pats realnames = let find_tomatch_tycon evdref env loc = function (* Try if some 'in I ...' is present and can be used as a constraint *) | Some (_,ind,realnal) -> - mk_tycon (inductive_template evdref env loc ind),Some (List.rev realnal) + mk_tycon (EConstr.of_constr (inductive_template evdref env loc ind)),Some (List.rev realnal) | None -> empty_tycon,None @@ -1211,7 +1211,7 @@ let rec generalize_problem names pb = function (* No more patterns: typing the right-hand side of equations *) let build_leaf pb = let rhs = extract_rhs pb in - let j = pb.typing_function (mk_tycon pb.pred) rhs.rhs_env pb.evdref rhs.it in + let j = pb.typing_function (mk_tycon (EConstr.of_constr pb.pred)) rhs.rhs_env pb.evdref rhs.it in j_nf_evar !(pb.evdref) j (* Build the sub-pattern-matching problem for a given branch "C x1..xn as x" *) @@ -1972,7 +1972,7 @@ let prepare_predicate loc typing_fun env sigma tomatchs arsign tycon pred = let envar = List.fold_right push_rel_context arsign env in let sigma, newt = new_sort_variable univ_flexible_alg sigma in let evdref = ref sigma in - let predcclj = typing_fun (mk_tycon (mkSort newt)) envar evdref rtntyp in + let predcclj = typing_fun (mk_tycon (EConstr.mkSort newt)) envar evdref rtntyp in let sigma = !evdref in let predccl = (j_nf_evar sigma predcclj).uj_val in [sigma, predccl] @@ -2238,7 +2238,7 @@ let constrs_of_pats typing_fun env evdref eqns tomatchs sign neqs arity = eqs_rels @ neqs_rels @ rhs_rels', arity in let rhs_env = push_rel_context rhs_rels' env in - let j = typing_fun (mk_tycon tycon) rhs_env eqn.rhs.it in + let j = typing_fun (mk_tycon (EConstr.of_constr tycon)) rhs_env eqn.rhs.it in let bbody = it_mkLambda_or_LetIn j.uj_val rhs_rels' and btype = it_mkProd_or_LetIn j.uj_type rhs_rels' in let _btype = evd_comb1 (Typing.type_of env) evdref bbody in diff --git a/pretyping/coercion.ml b/pretyping/coercion.ml index 6a7f90463..c5418d22e 100644 --- a/pretyping/coercion.ml +++ b/pretyping/coercion.ml @@ -244,8 +244,9 @@ and coerce loc env evdref (x : Term.constr) (y : Term.constr) match kind_of_term c with | Lambda (n, t, t') -> c, t' | Evar (k, args) -> - let (evs, t) = Evardefine.define_evar_as_lambda env !evdref (k,args) in + let (evs, t) = Evardefine.define_evar_as_lambda env !evdref (k, Array.map EConstr.of_constr args) in evdref := evs; + let t = EConstr.Unsafe.to_constr t in let (n, dom, rng) = destLambda t in let dom = whd_evar !evdref dom in if isEvar dom then @@ -374,11 +375,11 @@ let apply_coercion env sigma p hj typ_cl = (* Try to coerce to a funclass; raise NoCoercion if not possible *) let inh_app_fun_core env evd j = let t = whd_all env evd (EConstr.of_constr j.uj_type) in - match kind_of_term t with + match EConstr.kind evd (EConstr.of_constr t) with | Prod (_,_,_) -> (evd,j) | Evar ev -> let (evd',t) = Evardefine.define_evar_as_product evd ev in - (evd',{ uj_val = j.uj_val; uj_type = t }) + (evd',{ uj_val = j.uj_val; uj_type = EConstr.Unsafe.to_constr t }) | _ -> try let t,p = lookup_path_to_fun_from env evd j.uj_type in @@ -415,9 +416,9 @@ let inh_tosort_force loc env evd j = let inh_coerce_to_sort loc env evd j = let typ = whd_all env evd (EConstr.of_constr j.uj_type) in - match kind_of_term typ with + match EConstr.kind evd (EConstr.of_constr typ) with | Sort s -> (evd,{ utj_val = j.uj_val; utj_type = s }) - | Evar ev when not (is_defined evd (fst ev)) -> + | Evar ev -> let (evd',s) = Evardefine.define_evar_as_sort env evd ev in (evd',{ utj_val = j.uj_val; utj_type = s }) | _ -> diff --git a/pretyping/evarconv.ml b/pretyping/evarconv.ml index 47db71cc6..4540af28b 100644 --- a/pretyping/evarconv.ml +++ b/pretyping/evarconv.ml @@ -503,7 +503,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty let i,tF = if isRel tR || isVar tR then (* Optimization so as to generate candidates *) - let i,ev = evar_absorb_arguments env i ev (List.map EConstr.Unsafe.to_constr lF) in + let i,ev = evar_absorb_arguments env i (fst ev, Array.map EConstr.of_constr (snd ev)) lF in i,mkEvar ev else i,zip evd apprF in diff --git a/pretyping/evardefine.ml b/pretyping/evardefine.ml index 3982edd1c..8026ff3e4 100644 --- a/pretyping/evardefine.ml +++ b/pretyping/evardefine.ml @@ -61,13 +61,13 @@ type val_constraint = constr option let empty_tycon = None (* Builds a type constraint *) -let mk_tycon ty = Some ty +let mk_tycon ty = Some (EConstr.Unsafe.to_constr ty) (* Constrains the value of a type *) let empty_valcon = None (* Builds a value constraint *) -let mk_valcon c = Some c +let mk_valcon c = Some (EConstr.Unsafe.to_constr c) let idx = Namegen.default_dependent_ident @@ -75,11 +75,12 @@ let idx = Namegen.default_dependent_ident let define_pure_evar_as_product evd evk = let open Context.Named.Declaration in + let open EConstr in let evi = Evd.find_undefined evd evk in let evenv = evar_env evi in let id = next_ident_away idx (ids_of_named_context (evar_context evi)) in let concl = Reductionops.whd_all evenv evd (EConstr.of_constr evi.evar_concl) in - let s = destSort concl in + let s = destSort evd (EConstr.of_constr concl) in let evd1,(dom,u1) = let evd = Sigma.Unsafe.of_evar_map evd in let Sigma (e, evd1, _) = new_type_evar evenv evd univ_flexible_alg ~filter:(evar_filter evi) in @@ -103,20 +104,21 @@ let define_pure_evar_as_product evd evk = let evd3 = Evd.set_leq_sort evenv evd3 (Type prods) s in evd3, rng in - let prod = mkProd (Name id, dom, subst_var id rng) in - let evd3 = Evd.define evk prod evd2 in + let prod = mkProd (Name id, EConstr.of_constr dom, EConstr.of_constr (subst_var id rng)) in + let evd3 = Evd.define evk (EConstr.Unsafe.to_constr prod) evd2 in evd3,prod (* Refine an applied evar to a product and returns its instantiation *) let define_evar_as_product evd (evk,args) = + let open EConstr in let evd,prod = define_pure_evar_as_product evd evk in (* Quick way to compute the instantiation of evk with args *) - let na,dom,rng = destProd prod in - let evdom = mkEvar (fst (destEvar dom), args) in - let evrngargs = Array.cons (mkRel 1) (Array.map (lift 1) args) in - let evrng = mkEvar (fst (destEvar rng), evrngargs) in - evd,mkProd (na, evdom, evrng) + let na,dom,rng = destProd evd prod in + let evdom = mkEvar (fst (destEvar evd dom), args) in + let evrngargs = Array.cons (mkRel 1) (Array.map (Vars.lift 1) args) in + let evrng = mkEvar (fst (destEvar evd rng), evrngargs) in + evd, mkProd (na, evdom, evrng) (* Refine an evar with an abstraction @@ -129,38 +131,42 @@ let define_evar_as_product evd (evk,args) = let define_pure_evar_as_lambda env evd evk = let open Context.Named.Declaration in + let open EConstr in let evi = Evd.find_undefined evd evk in let evenv = evar_env evi in - let typ = Reductionops.whd_all evenv evd (EConstr.of_constr (evar_concl evi)) in - let evd1,(na,dom,rng) = match kind_of_term typ with + let typ = EConstr.of_constr (Reductionops.whd_all evenv evd (EConstr.of_constr (evar_concl evi))) in + let evd1,(na,dom,rng) = match EConstr.kind evd typ with | Prod (na,dom,rng) -> (evd,(na,dom,rng)) - | Evar ev' -> let evd,typ = define_evar_as_product evd ev' in evd,destProd typ + | Evar ev' -> let evd,typ = define_evar_as_product evd ev' in evd,destProd evd typ | _ -> error_not_product env evd typ in let avoid = ids_of_named_context (evar_context evi) in + let dom = EConstr.Unsafe.to_constr dom in let id = next_name_away_with_default_using_types "x" na avoid (Reductionops.whd_evar evd dom) in let newenv = push_named (LocalAssum (id, dom)) evenv in let filter = Filter.extend 1 (evar_filter evi) in let src = evar_source evk evd1 in - let evd2,body = new_evar_unsafe newenv evd1 ~src (subst1 (mkVar id) rng) ~filter in - let lam = mkLambda (Name id, dom, subst_var id body) in - Evd.define evk lam evd2, lam + let evd2,body = new_evar_unsafe newenv evd1 ~src (EConstr.Unsafe.to_constr (Vars.subst1 (mkVar id) rng)) ~filter in + let lam = mkLambda (Name id, EConstr.of_constr dom, Vars.subst_var id (EConstr.of_constr body)) in + Evd.define evk (EConstr.Unsafe.to_constr lam) evd2, lam let define_evar_as_lambda env evd (evk,args) = + let open EConstr in let evd,lam = define_pure_evar_as_lambda env evd evk in (* Quick way to compute the instantiation of evk with args *) - let na,dom,body = destLambda lam in - let evbodyargs = Array.cons (mkRel 1) (Array.map (lift 1) args) in - let evbody = mkEvar (fst (destEvar body), evbodyargs) in - evd,mkLambda (na, dom, evbody) + let na,dom,body = destLambda evd lam in + let evbodyargs = Array.cons (mkRel 1) (Array.map (Vars.lift 1) args) in + let evbody = mkEvar (fst (destEvar evd body), evbodyargs) in + evd, mkLambda (na, dom, evbody) let rec evar_absorb_arguments env evd (evk,args as ev) = function - | [] -> evd,ev + | [] -> evd, (evk, Array.map EConstr.Unsafe.to_constr args) | a::l -> + let open EConstr in (* TODO: optimize and avoid introducing intermediate evars *) let evd,lam = define_pure_evar_as_lambda env evd evk in - let _,_,body = destLambda lam in - let evk = fst (destEvar body) in + let _,_,body = destLambda evd lam in + let evk = fst (destEvar evd body) in evar_absorb_arguments env evd (evk, Array.cons a args) l (* Refining an evar to a sort *) @@ -180,23 +186,24 @@ let define_evar_as_sort env evd (ev,args) = an evar instantiate it with the product of 2 new evars. *) let split_tycon loc env evd tycon = + let open EConstr in let rec real_split evd c = - let t = Reductionops.whd_all env evd (EConstr.of_constr c) in - match kind_of_term t with + let t = Reductionops.whd_all env evd c in + match EConstr.kind evd (EConstr.of_constr t) with | Prod (na,dom,rng) -> evd, (na, dom, rng) | Evar ev (* ev is undefined because of whd_all *) -> let (evd',prod) = define_evar_as_product evd ev in - let (_,dom,rng) = destProd prod in + let (_,dom,rng) = destProd evd prod in evd',(Anonymous, dom, rng) - | App (c,args) when isEvar c -> - let (evd',lam) = define_evar_as_lambda env evd (destEvar c) in + | App (c,args) when isEvar evd c -> + let (evd',lam) = define_evar_as_lambda env evd (destEvar evd c) in real_split evd' (mkApp (lam,args)) | _ -> error_not_product ~loc env evd c in match tycon with | None -> evd,(Anonymous,None,None) | Some c -> - let evd', (n, dom, rng) = real_split evd c in + let evd', (n, dom, rng) = real_split evd (EConstr.of_constr c) in evd', (n, mk_tycon dom, mk_tycon rng) let valcon_of_tycon x = x diff --git a/pretyping/evardefine.mli b/pretyping/evardefine.mli index 07b0e69d9..f6d0efba6 100644 --- a/pretyping/evardefine.mli +++ b/pretyping/evardefine.mli @@ -18,15 +18,15 @@ type type_constraint = types option type val_constraint = constr option val empty_tycon : type_constraint -val mk_tycon : constr -> type_constraint +val mk_tycon : EConstr.constr -> type_constraint val empty_valcon : val_constraint -val mk_valcon : constr -> val_constraint +val mk_valcon : EConstr.constr -> val_constraint (** Instantiate an evar by as many lambda's as needed so that its arguments are moved to the evar substitution (i.e. turn [?x[vars1:=args1] args] into [?y[vars1:=args1,vars:=args]] with [vars1 |- ?x:=\vars.?y[vars1:=vars1,vars:=vars]] *) -val evar_absorb_arguments : env -> evar_map -> existential -> constr list -> +val evar_absorb_arguments : env -> evar_map -> EConstr.existential -> EConstr.constr list -> evar_map * existential val split_tycon : @@ -36,9 +36,9 @@ val split_tycon : val valcon_of_tycon : type_constraint -> val_constraint val lift_tycon : int -> type_constraint -> type_constraint -val define_evar_as_product : evar_map -> existential -> evar_map * types -val define_evar_as_lambda : env -> evar_map -> existential -> evar_map * types -val define_evar_as_sort : env -> evar_map -> existential -> evar_map * sorts +val define_evar_as_product : evar_map -> EConstr.existential -> evar_map * EConstr.types +val define_evar_as_lambda : env -> evar_map -> EConstr.existential -> evar_map * EConstr.types +val define_evar_as_sort : env -> evar_map -> EConstr.existential -> evar_map * sorts (** {6 debug pretty-printer:} *) diff --git a/pretyping/pretype_errors.ml b/pretyping/pretype_errors.ml index f8f6d44bf..f28fb84ba 100644 --- a/pretyping/pretype_errors.ml +++ b/pretyping/pretype_errors.ml @@ -44,14 +44,14 @@ type pretype_error = | CannotUnifyBindingType of constr * constr | CannotGeneralize of constr | NoOccurrenceFound of constr * Id.t option - | CannotFindWellTypedAbstraction of constr * constr list * (env * type_error) option + | CannotFindWellTypedAbstraction of constr * EConstr.constr list * (env * type_error) option | WrongAbstractionType of Name.t * constr * types * types | AbstractionOverMeta of Name.t * Name.t | NonLinearUnification of Name.t * constr (* Pretyping *) | VarNotFound of Id.t | UnexpectedType of constr * constr - | NotProduct of constr + | NotProduct of EConstr.constr | TypingError of type_error | CannotUnifyOccurrences of subterm_unification_error | UnsatisfiableConstraints of diff --git a/pretyping/pretype_errors.mli b/pretyping/pretype_errors.mli index b015add79..8a6d8b6b3 100644 --- a/pretyping/pretype_errors.mli +++ b/pretyping/pretype_errors.mli @@ -45,14 +45,14 @@ type pretype_error = | CannotUnifyBindingType of constr * constr | CannotGeneralize of constr | NoOccurrenceFound of constr * Id.t option - | CannotFindWellTypedAbstraction of constr * constr list * (env * type_error) option + | CannotFindWellTypedAbstraction of constr * EConstr.constr list * (env * type_error) option | WrongAbstractionType of Name.t * constr * types * types | AbstractionOverMeta of Name.t * Name.t | NonLinearUnification of Name.t * constr (** Pretyping *) | VarNotFound of Id.t | UnexpectedType of constr * constr - | NotProduct of constr + | NotProduct of EConstr.constr | TypingError of type_error | CannotUnifyOccurrences of subterm_unification_error | UnsatisfiableConstraints of @@ -110,7 +110,7 @@ val error_cannot_unify : ?loc:Loc.t -> env -> Evd.evar_map -> val error_cannot_unify_local : env -> Evd.evar_map -> constr * constr * constr -> 'b val error_cannot_find_well_typed_abstraction : env -> Evd.evar_map -> - constr -> constr list -> (env * type_error) option -> 'b + constr -> EConstr.constr list -> (env * type_error) option -> 'b val error_wrong_abstraction_type : env -> Evd.evar_map -> Name.t -> constr -> types -> types -> 'b @@ -132,7 +132,7 @@ val error_unexpected_type : ?loc:Loc.t -> env -> Evd.evar_map -> constr -> constr -> 'b val error_not_product : - ?loc:Loc.t -> env -> Evd.evar_map -> constr -> 'b + ?loc:Loc.t -> env -> Evd.evar_map -> EConstr.constr -> 'b (** {6 Error in conversion from AST to glob_constr } *) diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml index 3c48c42df..b689bb7c7 100644 --- a/pretyping/pretyping.ml +++ b/pretyping/pretyping.ml @@ -606,7 +606,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre type_bl (push_rel dcl env) (Context.Rel.add dcl' ctxt) bl | (na,bk,Some bd,ty)::bl -> let ty' = pretype_type empty_valcon env evdref lvar ty in - let bd' = pretype (mk_tycon ty'.utj_val) env evdref lvar bd in + let bd' = pretype (mk_tycon (EConstr.of_constr ty'.utj_val)) env evdref lvar bd in let dcl = LocalDef (na, bd'.uj_val, ty'.utj_val) in let dcl' = LocalDef (ltac_interp_name lvar na, bd'.uj_val, ty'.utj_val) in type_bl (push_rel dcl env) (Context.Rel.add dcl' ctxt) bl in @@ -640,7 +640,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre decompose_prod_n_assum (Context.Rel.length ctxt) (lift nbfix ftys.(i)) in let nenv = push_rel_context ctxt newenv in - let j = pretype (mk_tycon ty) nenv evdref lvar def in + let j = pretype (mk_tycon (EConstr.of_constr ty)) nenv evdref lvar def in { uj_val = it_mkLambda_or_LetIn j.uj_val ctxt; uj_type = it_mkProd_or_LetIn j.uj_type ctxt }) ctxtv vdef in @@ -815,7 +815,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre match c1 with | GCast (loc, c, CastConv t) -> let tj = pretype_type empty_valcon env evdref lvar t in - pretype (mk_tycon tj.utj_val) env evdref lvar c + pretype (mk_tycon (EConstr.of_constr tj.utj_val)) env evdref lvar c | _ -> pretype empty_tycon env evdref lvar c1 in let t = evd_comb1 (Evarsolve.refresh_universes @@ -895,7 +895,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre @[build_dependent_constructor cs] in let lp = lift cs.cs_nargs p in let fty = hnf_lam_applist env.ExtraEnv.env !evdref (EConstr.of_constr lp) (List.map EConstr.of_constr inst) in - let fj = pretype (mk_tycon fty) env_f evdref lvar d in + let fj = pretype (mk_tycon (EConstr.of_constr fty)) env_f evdref lvar d in let v = let ind,_ = dest_ind_family indf in Typing.check_allowed_sort env.ExtraEnv.env !evdref ind cj.uj_val p; @@ -973,7 +973,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre cs.cs_args in let env_c = push_rel_context csgn env in - let bj = pretype (mk_tycon pi) env_c evdref lvar b in + let bj = pretype (mk_tycon (EConstr.of_constr pi)) env_c evdref lvar b in it_mkLambda_or_LetIn bj.uj_val cs.cs_args in let b1 = f cstrs.(0) b1 in let b2 = f cstrs.(1) b2 in @@ -1028,7 +1028,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre (ConversionFailed (env.ExtraEnv.env,cty,tval)) end | _ -> - pretype (mk_tycon tval) env evdref lvar c, tval + pretype (mk_tycon (EConstr.of_constr tval)) env evdref lvar c, tval in let v = mkCast (cj.uj_val, k, tval) in { uj_val = v; uj_type = tval } @@ -1041,7 +1041,7 @@ and pretype_instance k0 resolve_tc env evdref lvar loc hyps evk update = let c, update = try let c = List.assoc id update in - let c = pretype k0 resolve_tc (mk_tycon t) env evdref lvar c in + let c = pretype k0 resolve_tc (mk_tycon (EConstr.of_constr t)) env evdref lvar c in c.uj_val, List.remove_assoc id update with Not_found -> try @@ -1068,9 +1068,9 @@ and pretype_type k0 resolve_tc valcon (env : ExtraEnv.t) evdref lvar = function let s = let sigma = !evdref in let t = Retyping.get_type_of env.ExtraEnv.env sigma (EConstr.of_constr v) in - match kind_of_term (whd_all env.ExtraEnv.env sigma (EConstr.of_constr t)) with + match EConstr.kind sigma (EConstr.of_constr (whd_all env.ExtraEnv.env sigma (EConstr.of_constr t))) with | Sort s -> s - | Evar ev when is_Type (existential_type sigma ev) -> + | Evar ev when is_Type (existential_type sigma (fst ev, Array.map EConstr.Unsafe.to_constr (snd ev))) -> evd_comb1 (define_evar_as_sort env.ExtraEnv.env) evdref ev | _ -> anomaly (Pp.str "Found a type constraint which is not a type") in @@ -1101,7 +1101,7 @@ let ise_pretype_gen flags env sigma lvar kind c = | WithoutTypeConstraint -> (pretype k0 flags.use_typeclasses empty_tycon env evdref lvar c).uj_val | OfType exptyp -> - (pretype k0 flags.use_typeclasses (mk_tycon exptyp) env evdref lvar c).uj_val + (pretype k0 flags.use_typeclasses (mk_tycon (EConstr.of_constr exptyp)) env evdref lvar c).uj_val | IsType -> (pretype_type k0 flags.use_typeclasses empty_valcon env evdref lvar c).utj_val in diff --git a/pretyping/typing.ml b/pretyping/typing.ml index acfe05f24..db31541cd 100644 --- a/pretyping/typing.ml +++ b/pretyping/typing.ml @@ -36,7 +36,7 @@ let inductive_type_knowing_parameters env (ind,u) jl = Inductive.type_of_inductive_knowing_parameters env (mspec,u) paramstyp let e_type_judgment env evdref j = - match kind_of_term (whd_all env !evdref (EConstr.of_constr j.uj_type)) with + match EConstr.kind !evdref (EConstr.of_constr (whd_all env !evdref (EConstr.of_constr j.uj_type))) with | Sort s -> {utj_val = j.uj_val; utj_type = s } | Evar ev -> let (evd,s) = Evardefine.define_evar_as_sort env !evdref ev in @@ -49,26 +49,27 @@ let e_assumption_of_judgment env evdref j = error_assumption env j let e_judge_of_apply env evdref funj argjv = + let open EConstr in let rec apply_rec n typ = function | [] -> - { uj_val = mkApp (j_val funj, Array.map j_val argjv); - uj_type = typ } + { uj_val = Constr.mkApp (j_val funj, Array.map j_val argjv); + uj_type = EConstr.Unsafe.to_constr typ } | hj::restjl -> - match kind_of_term (whd_all env !evdref (EConstr.of_constr typ)) with + match EConstr.kind !evdref (EConstr.of_constr (whd_all env !evdref typ)) with | Prod (_,c1,c2) -> - if Evarconv.e_cumul env evdref hj.uj_type c1 then - apply_rec (n+1) (subst1 hj.uj_val c2) restjl + if Evarconv.e_cumul env evdref hj.uj_type (EConstr.Unsafe.to_constr c1) then + apply_rec (n+1) (Vars.subst1 (EConstr.of_constr hj.uj_val) c2) restjl else - error_cant_apply_bad_type env (n,c1, hj.uj_type) funj argjv + error_cant_apply_bad_type env (n, EConstr.Unsafe.to_constr c1, hj.uj_type) funj argjv | Evar ev -> let (evd',t) = Evardefine.define_evar_as_product !evdref ev in evdref := evd'; - let (_,_,c2) = destProd t in - apply_rec (n+1) (subst1 hj.uj_val c2) restjl + let (_,_,c2) = destProd evd' t in + apply_rec (n+1) (Vars.subst1 (EConstr.of_constr hj.uj_val) c2) restjl | _ -> error_cant_apply_not_functional env funj argjv in - apply_rec 1 funj.uj_type (Array.to_list argjv) + apply_rec 1 (EConstr.of_constr funj.uj_type) (Array.to_list argjv) let e_check_branch_types env evdref (ind,u) cj (lfj,explft) = if not (Int.equal (Array.length lfj) (Array.length explft)) then diff --git a/pretyping/unification.ml b/pretyping/unification.ml index 8865e69ef..f282ec4f1 100644 --- a/pretyping/unification.ml +++ b/pretyping/unification.ml @@ -105,26 +105,27 @@ let abstract_list_all env evd typ c l = try Typing.type_of env evd p with | UserError _ -> - error_cannot_find_well_typed_abstraction env evd p l None + error_cannot_find_well_typed_abstraction env evd p (List.map EConstr.of_constr l) None | Type_errors.TypeError (env',x) -> - error_cannot_find_well_typed_abstraction env evd p l (Some (env',x)) in + error_cannot_find_well_typed_abstraction env evd p (List.map EConstr.of_constr l) (Some (env',x)) in evd,(p,typp) let set_occurrences_of_last_arg args = Some AllOccurrences :: List.tl (Array.map_to_list (fun _ -> None) args) let abstract_list_all_with_dependencies env evd typ c l = + let open EConstr in let evd = Sigma.Unsafe.of_evar_map evd in let Sigma (ev, evd, _) = new_evar env evd typ in let evd = Sigma.to_evar_map evd in - let evd,ev' = evar_absorb_arguments env evd (destEvar ev) l in + let evd,ev' = evar_absorb_arguments env evd (destEvar evd (EConstr.of_constr ev)) l in let n = List.length l in let argoccs = set_occurrences_of_last_arg (Array.sub (snd ev') 0 n) in let evd,b = Evarconv.second_order_matching empty_transparent_state env evd ev' argoccs c in if b then - let p = nf_evar evd (existential_value evd (destEvar ev)) in + let p = nf_evar evd ev in evd, p else error_cannot_find_well_typed_abstraction env evd (nf_evar evd c) l None @@ -1899,7 +1900,7 @@ let secondOrderAbstraction env evd flags typ (p, oplist) = let secondOrderDependentAbstraction env evd flags typ (p, oplist) = let typp = Typing.meta_type evd p in - let evd, pred = abstract_list_all_with_dependencies env evd typp typ oplist in + let evd, pred = abstract_list_all_with_dependencies env evd typp typ (List.map EConstr.of_constr oplist) in w_merge env false flags.merge_unify_flags (evd,[p,pred,(Conv,TypeProcessed)],[]) diff --git a/toplevel/himsg.ml b/toplevel/himsg.ml index 4cffbee82..c09bf59ce 100644 --- a/toplevel/himsg.ml +++ b/toplevel/himsg.ml @@ -405,7 +405,7 @@ let explain_unexpected_type env sigma actual_type expected_type = str "where" ++ spc () ++ prexp ++ str " was expected." let explain_not_product env sigma c = - let c = Evarutil.nf_evar sigma c in + let c = EConstr.to_constr sigma c in let pr = pr_lconstr_env env sigma c in str "The type of this term is a product" ++ spc () ++ str "while it is expected to be" ++ @@ -654,7 +654,7 @@ let explain_cannot_unify_binding_type env sigma m n = let explain_cannot_find_well_typed_abstraction env sigma p l e = str "Abstracting over the " ++ str (String.plural (List.length l) "term") ++ spc () ++ - hov 0 (pr_enum (pr_lconstr_env env sigma) l) ++ spc () ++ + hov 0 (pr_enum (fun c -> pr_lconstr_env env sigma (EConstr.to_constr sigma c)) l) ++ spc () ++ str "leads to a term" ++ spc () ++ pr_lconstr_goal_style_env env sigma p ++ spc () ++ str "which is ill-typed." ++ (match e with None -> mt () | Some e -> fnl () ++ str "Reason is: " ++ e) -- cgit v1.2.3 From b7fd585b89ac5e0b7770f52739c33fe179f2eed8 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sat, 5 Nov 2016 21:36:40 +0100 Subject: Evarsolve API using EConstr. --- engine/evarutil.mli | 2 +- plugins/cc/cctac.ml | 6 +- pretyping/cases.ml | 12 +- pretyping/evarconv.ml | 85 ++++---- pretyping/evarsolve.ml | 475 +++++++++++++++++++++++-------------------- pretyping/evarsolve.mli | 11 +- pretyping/pretype_errors.ml | 8 +- pretyping/pretype_errors.mli | 8 +- pretyping/pretyping.ml | 10 +- pretyping/typing.ml | 4 +- pretyping/unification.ml | 21 +- tactics/tactics.ml | 6 +- toplevel/himsg.ml | 22 +- 13 files changed, 364 insertions(+), 306 deletions(-) diff --git a/engine/evarutil.mli b/engine/evarutil.mli index 7fdc7aac7..a2e2a07dd 100644 --- a/engine/evarutil.mli +++ b/engine/evarutil.mli @@ -76,7 +76,7 @@ val new_evar_instance : ?principal:bool -> constr list -> (constr, 'r) Sigma.sigma -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 diff --git a/plugins/cc/cctac.ml b/plugins/cc/cctac.ml index 425bb2d6f..36a96fdb5 100644 --- a/plugins/cc/cctac.ml +++ b/plugins/cc/cctac.ml @@ -259,7 +259,7 @@ let refresh_universes ty k = Proofview.Goal.enter { enter = begin fun gl -> let env = Proofview.Goal.env gl in let evm = Tacmach.New.project gl in - let evm, ty = refresh_type env evm ty in + let evm, ty = refresh_type env evm (EConstr.of_constr ty) in Tacticals.New.tclTHEN (Proofview.V82.tactic (Refiner.tclEVARS evm)) (k ty) end } @@ -376,7 +376,7 @@ let discriminate_tac (cstr,u as cstru) p = let identity = Universes.constr_of_global (Lazy.force _I) in let trivial = Universes.constr_of_global (Lazy.force _True) in let evm = Tacmach.New.project gl in - let evm, intype = refresh_type env evm (Tacmach.New.pf_unsafe_type_of gl t1) in + let evm, intype = refresh_type env evm (EConstr.of_constr (Tacmach.New.pf_unsafe_type_of gl t1)) in let evm, outtype = Evd.new_sort_variable Evd.univ_flexible evm in let outtype = mkSort outtype in let pred = mkLambda(Name xid,outtype,mkRel 1) in @@ -481,7 +481,7 @@ let mk_eq f c1 c2 k = Proofview.Goal.enter { enter = begin fun gl -> let open Tacmach.New in let evm, ty = pf_apply type_of gl c1 in - let evm, ty = Evarsolve.refresh_universes (Some false) (pf_env gl) evm ty in + let evm, ty = Evarsolve.refresh_universes (Some false) (pf_env gl) evm (EConstr.of_constr ty) in let term = mkApp (fc, [| ty; c1; c2 |]) in let evm, _ = type_of (pf_env gl) evm term in Tacticals.New.tclTHEN (Proofview.V82.tactic (Refiner.tclEVARS evm)) diff --git a/pretyping/cases.ml b/pretyping/cases.ml index 915cd261d..a68daf4e5 100644 --- a/pretyping/cases.ml +++ b/pretyping/cases.ml @@ -78,6 +78,9 @@ let list_try_compile f l = let force_name = let nx = Name default_dependent_ident in function Anonymous -> nx | na -> na +let to_conv_fun f = (); fun env sigma pb c1 c2 -> + f env sigma pb (EConstr.Unsafe.to_constr c1) (EConstr.Unsafe.to_constr c2) + (************************************************************************) (* Pattern-matching compilation (Cases) *) (************************************************************************) @@ -1583,7 +1586,7 @@ let adjust_to_extended_env_and_remove_deps env extenv sigma subst t = | LocalAssum _ -> p in let p = traverse_local_defs p in let u = lift (n' - n) u in - try Some (p, u, expand_vars_in_term extenv sigma u) + try Some (p, u, EConstr.Unsafe.to_constr (expand_vars_in_term extenv sigma (EConstr.of_constr u))) (* pedrot: does this really happen to raise [Failure _]? *) with Failure _ -> None in let subst0 = List.map_filter map subst in @@ -1628,14 +1631,15 @@ let abstract_tycon loc env evdref subst tycon extenv t = | Rel n when is_local_def (lookup_rel n env) -> t | Evar ev -> let ty = get_type_of env !evdref (EConstr.of_constr t) in - let ty = Evarutil.evd_comb1 (refresh_universes (Some false) env) evdref ty in + let ty = Evarutil.evd_comb1 (refresh_universes (Some false) env) evdref (EConstr.of_constr ty) in let inst = List.map_i (fun i _ -> try list_assoc_in_triple i subst0 with Not_found -> mkRel i) 1 (rel_context env) in let ev' = e_new_evar env evdref ~src ty in - begin match solve_simple_eqn (evar_conv_x full_transparent_state) env !evdref (None,ev,EConstr.of_constr (substl inst ev')) with + let ev = (fst ev, Array.map EConstr.of_constr (snd ev)) in + begin match solve_simple_eqn (to_conv_fun (evar_conv_x full_transparent_state)) env !evdref (None,ev,EConstr.of_constr (substl inst ev')) with | Success evd -> evdref := evd | UnifFailure _ -> assert false end; @@ -1650,7 +1654,7 @@ let abstract_tycon loc env evdref subst tycon extenv t = let vl = List.map pi1 good in let ty = let ty = get_type_of env !evdref (EConstr.of_constr t) in - Evarutil.evd_comb1 (refresh_universes (Some false) env) evdref ty + Evarutil.evd_comb1 (refresh_universes (Some false) env) evdref (EConstr.of_constr ty) in let ty = lift (-k) (aux x ty) in let depvl = free_rels !evdref (EConstr.of_constr ty) in diff --git a/pretyping/evarconv.ml b/pretyping/evarconv.ml index 4540af28b..8f3f671ab 100644 --- a/pretyping/evarconv.ml +++ b/pretyping/evarconv.ml @@ -42,6 +42,9 @@ let _ = Goptions.declare_bool_option { Goptions.optwrite = (fun a -> debug_unification:=a); } +let to_conv_fun f = (); fun env sigma pb c1 c2 -> + f env sigma pb (EConstr.Unsafe.to_constr c1) (EConstr.Unsafe.to_constr c2) + let unfold_projection env evd ts p c = let cst = Projection.constant p in if is_transparent_constant ts cst then @@ -341,7 +344,7 @@ let rec evar_conv_x ts env evd pbty term1 term2 = env evd term1 term2 in if b then Success evd - else UnifFailure (evd, ConversionFailed (env,term1,term2)) + else UnifFailure (evd, ConversionFailed (env,EConstr.of_constr term1,EConstr.of_constr term2)) with Univ.UniverseInconsistency e -> UnifFailure (evd, UnifUnivInconsistency e) in match e with @@ -361,15 +364,15 @@ let rec evar_conv_x ts env evd pbty term1 term2 = (whd_nored_state evd (EConstr.of_constr term1,Stack.empty), Cst_stack.empty) (whd_nored_state evd (EConstr.of_constr term2,Stack.empty), Cst_stack.empty) in - begin match kind_of_term term1, kind_of_term term2 with + begin match EConstr.kind evd (EConstr.of_constr term1), EConstr.kind evd (EConstr.of_constr term2) with | Evar ev, _ when Evd.is_undefined evd (fst ev) -> - (match solve_simple_eqn (evar_conv_x ts) env evd + (match solve_simple_eqn (to_conv_fun (evar_conv_x ts)) env evd (position_problem true pbty,ev, EConstr.of_constr term2) with | UnifFailure (_,OccurCheck _) -> (* Eta-expansion might apply *) default () | x -> x) | _, Evar ev when Evd.is_undefined evd (fst ev) -> - (match solve_simple_eqn (evar_conv_x ts) env evd + (match solve_simple_eqn (to_conv_fun (evar_conv_x ts)) env evd (position_problem false pbty,ev, EConstr.of_constr term1) with | UnifFailure (_, OccurCheck _) -> (* Eta-expansion might apply *) default () @@ -383,13 +386,12 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty UnifFailure (i, NotSameHead) in let miller_pfenning on_left fallback ev lF tM evd = - let lF = List.map EConstr.Unsafe.to_constr lF in match is_unification_pattern_evar env evd ev lF tM with | None -> fallback () | Some l1' -> (* Miller-Pfenning's patterns unification *) - let t2 = nf_evar evd tM in + let t2 = EConstr.of_constr (nf_evar evd (EConstr.Unsafe.to_constr tM)) (** FIXME *) in let t2 = solve_pattern_eqn env evd l1' t2 in - solve_simple_eqn (evar_conv_x ts) env evd + solve_simple_eqn (to_conv_fun (evar_conv_x ts)) env evd (position_problem on_left pbty,ev, EConstr.of_constr t2) in let consume_stack on_left (termF,skF) (termO,skO) evd = @@ -441,7 +443,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty (fun () -> if not_only_app then (* Postpone the use of an heuristic *) switch (fun x y -> Success (add_conv_pb (pbty,env,x,y) i)) (zip evd apprF) tM else quick_fail i) - ev lF tM i + ev lF (EConstr.of_constr tM) i and consume (termF,skF as apprF) (termM,skM as apprM) i = if not (Stack.is_empty skF && Stack.is_empty skM) then consume_stack on_left apprF apprM i @@ -510,8 +512,8 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty switch (fun x y -> Success (add_conv_pb (pbty,env,x,y) i)) tF tR else - UnifFailure (evd,OccurCheck (fst ev,tR)))]) - ev lF tR evd + UnifFailure (evd,OccurCheck (fst ev,EConstr.of_constr tR)))]) + (fst ev, Array.map EConstr.of_constr (snd ev)) lF (EConstr.of_constr tR) evd in let app_empty = match sk1, sk2 with [], [] -> true | _ -> false in (* Evar must be undefined since we have flushed evars *) @@ -529,33 +531,33 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty | None, Success i' -> (* We do have sk1[] = sk2[]: we now unify ?ev1 and ?ev2 *) (* Note that ?ev1 and ?ev2, may have been instantiated in the meantime *) - let ev1' = whd_evar i' (mkEvar ev1) in - if isEvar ev1' then - solve_simple_eqn (evar_conv_x ts) env i' - (position_problem true pbty,destEvar ev1', term2) + let ev1' = EConstr.of_constr (whd_evar i' (mkEvar ev1)) in + if EConstr.isEvar i' ev1' then + solve_simple_eqn (to_conv_fun (evar_conv_x ts)) env i' + (position_problem true pbty,EConstr.destEvar i' ev1', term2) else evar_eqappr_x ts env evd pbty - ((EConstr.of_constr ev1', sk1), csts1) ((term2, sk2), csts2) + ((ev1', sk1), csts1) ((term2, sk2), csts2) | Some (r,[]), Success i' -> (* We have sk1'[] = sk2[] for some sk1' s.t. sk1[]=sk1'[r[]] *) (* we now unify r[?ev1] and ?ev2 *) - let ev2' = whd_evar i' (mkEvar ev2) in - if isEvar ev2' then - solve_simple_eqn (evar_conv_x ts) env i' - (position_problem false pbty,destEvar ev2',Stack.zip evd (term1,r)) + let ev2' = EConstr.of_constr (whd_evar i' (mkEvar ev2)) in + if EConstr.isEvar i' ev2' then + solve_simple_eqn (to_conv_fun (evar_conv_x ts)) env i' + (position_problem false pbty,EConstr.destEvar i' ev2',Stack.zip evd (term1,r)) else evar_eqappr_x ts env evd pbty - ((EConstr.of_constr ev2', sk1), csts1) ((term2, sk2), csts2) + ((ev2', sk1), csts1) ((term2, sk2), csts2) | Some ([],r), Success i' -> (* Symmetrically *) (* We have sk1[] = sk2'[] for some sk2' s.t. sk2[]=sk2'[r[]] *) (* we now unify ?ev1 and r[?ev2] *) - let ev1' = whd_evar i' (mkEvar ev1) in - if isEvar ev1' then - solve_simple_eqn (evar_conv_x ts) env i' - (position_problem true pbty,destEvar ev1',Stack.zip evd (term2,r)) + let ev1' = EConstr.of_constr (whd_evar i' (mkEvar ev1)) in + if EConstr.isEvar i' ev1' then + solve_simple_eqn (to_conv_fun (evar_conv_x ts)) env i' + (position_problem true pbty,EConstr.destEvar i' ev1',Stack.zip evd (term2,r)) else evar_eqappr_x ts env evd pbty - ((EConstr.of_constr ev1', sk1), csts1) ((term2, sk2), csts2) + ((ev1', sk1), csts1) ((term2, sk2), csts2) | None, (UnifFailure _ as x) -> (* sk1 and sk2 have no common outer part *) if Stack.not_purely_applicative sk2 then @@ -590,9 +592,9 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty if Evar.equal sp1 sp2 then match ise_stack2 false env i (evar_conv_x ts) sk1 sk2 with |None, Success i' -> - Success (solve_refl (fun env i pbty a1 a2 -> - is_success (evar_conv_x ts env i pbty a1 a2)) - env i' (position_problem true pbty) sp1 al1 al2) + Success (solve_refl (to_conv_fun (fun env i pbty a1 a2 -> + is_success (evar_conv_x ts env i pbty a1 a2))) + env i' (position_problem true pbty) sp1 (Array.map EConstr.of_constr al1) (Array.map EConstr.of_constr al2)) |_, (UnifFailure _ as x) -> x |Some _, _ -> UnifFailure (i,NotSameArgSize) else UnifFailure (i,NotSameHead) @@ -600,10 +602,10 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty ise_try evd [f1; f2] | Flexible ev1, MaybeFlexible v2 -> - flex_maybeflex true ev1 (appr1,csts1) (appr2,csts2) (EConstr.of_constr v2) + flex_maybeflex true (fst ev1, Array.map EConstr.of_constr (snd ev1)) (appr1,csts1) (appr2,csts2) (EConstr.of_constr v2) | MaybeFlexible v1, Flexible ev2 -> - flex_maybeflex false ev2 (appr2,csts2) (appr1,csts1) (EConstr.of_constr v1) + flex_maybeflex false (fst ev2, Array.map EConstr.of_constr (snd ev2)) (appr2,csts2) (appr1,csts1) (EConstr.of_constr v1) | MaybeFlexible v1, MaybeFlexible v2 -> begin match kind_of_term (EConstr.Unsafe.to_constr term1), kind_of_term (EConstr.Unsafe.to_constr term2) with @@ -964,7 +966,8 @@ let first_order_unification ts env evd (ev1,l1) (term2,l2) = if is_defined i (fst ev1) then evar_conv_x ts env i CONV t2 (mkEvar ev1) else - solve_simple_eqn ~choose:true (evar_conv_x ts) env i (None,ev1, EConstr.of_constr t2))] + let ev1 = (fst ev1, Array.map EConstr.of_constr (snd ev1)) in + solve_simple_eqn ~choose:true (to_conv_fun (evar_conv_x ts)) env i (None,ev1, EConstr.of_constr t2))] let choose_less_dependent_instance evk evd term args = let evi = Evd.find_undefined evd evk in @@ -1109,7 +1112,7 @@ let second_order_matching ts env_rhs evd (evk,args) argoccs rhs = match evar_conv_x ts env_evar evd CUMUL idty evty with | UnifFailure _ -> error "Cannot find an instance" | Success evd -> - match reconsider_conv_pbs (evar_conv_x ts) evd with + match reconsider_conv_pbs (to_conv_fun (evar_conv_x ts)) evd with | UnifFailure _ -> error "Cannot find an instance" | Success evd -> evd @@ -1123,8 +1126,8 @@ let second_order_matching ts env_rhs evd (evk,args) argoccs rhs = force_instantiation evd !evsref | [] -> let evd = - try Evarsolve.check_evar_instance evd evk rhs - (evar_conv_x full_transparent_state) + try Evarsolve.check_evar_instance evd evk (EConstr.of_constr rhs) + (to_conv_fun (evar_conv_x full_transparent_state)) with IllTypedInstance _ -> raise (TypingFailed evd) in Evd.define evk rhs evd @@ -1173,11 +1176,13 @@ let apply_conversion_problem_heuristic ts env evd pbty t1 t2 = UnifFailure (evd, CannotSolveConstraint ((pbty,env,t1,t2),reason))) | Evar (evk1,args1), Evar (evk2,args2) when Evar.equal evk1 evk2 -> let f env evd pbty x y = is_fconv ~reds:ts pbty env evd (EConstr.of_constr x) (EConstr.of_constr y) in - Success (solve_refl ~can_drop:true f env evd - (position_problem true pbty) evk1 args1 args2) + Success (solve_refl ~can_drop:true (to_conv_fun f) env evd + (position_problem true pbty) evk1 (Array.map EConstr.of_constr args1) (Array.map EConstr.of_constr args2)) | Evar ev1, Evar ev2 when app_empty -> + let ev1 = (fst ev1, Array.map EConstr.of_constr (snd ev1)) in + let ev2 = (fst ev2, Array.map EConstr.of_constr (snd ev2)) in Success (solve_evar_evar ~force:true - (evar_define (evar_conv_x ts) ~choose:true) (evar_conv_x ts) env evd + (evar_define (to_conv_fun (evar_conv_x ts)) ~choose:true) (to_conv_fun (evar_conv_x ts)) env evd (position_problem true pbty) ev1 ev2) | Evar ev1,_ when Array.length l1 <= Array.length l2 -> (* On "?n t1 .. tn = u u1 .. u(n+p)", try first-order unification *) @@ -1239,9 +1244,9 @@ let rec solve_unconstrained_evars_with_candidates ts evd = | a::l -> try let conv_algo = evar_conv_x ts in - let evd = check_evar_instance evd evk a conv_algo in + let evd = check_evar_instance evd evk (EConstr.of_constr a) (to_conv_fun conv_algo) in let evd = Evd.define evk a evd in - match reconsider_conv_pbs conv_algo evd with + match reconsider_conv_pbs (to_conv_fun conv_algo) evd with | Success evd -> solve_unconstrained_evars_with_candidates ts evd | UnifFailure _ -> aux l with @@ -1260,7 +1265,7 @@ let solve_unconstrained_impossible_cases env evd = let evd' = Evd.merge_context_set Evd.univ_flexible_alg ~loc evd' ctx in let ty = j_type j in let conv_algo = evar_conv_x full_transparent_state in - let evd' = check_evar_instance evd' evk ty conv_algo in + let evd' = check_evar_instance evd' evk (EConstr.of_constr ty) (to_conv_fun conv_algo) in Evd.define evk ty evd' | _ -> evd') evd evd diff --git a/pretyping/evarsolve.ml b/pretyping/evarsolve.ml index 86ef8f56f..8a22aed2f 100644 --- a/pretyping/evarsolve.ml +++ b/pretyping/evarsolve.ml @@ -22,13 +22,14 @@ open Pretype_errors open Sigma.Notations let normalize_evar evd ev = - match kind_of_term (whd_evar evd (mkEvar ev)) with + let open EConstr in + match EConstr.kind evd (mkEvar ev) with | Evar (evk,args) -> (evk,args) | _ -> assert false -let get_polymorphic_positions f = +let get_polymorphic_positions sigma f = let open Declarations in - match kind_of_term f with + match EConstr.kind sigma f with | Ind (ind, u) | Construct ((ind, _), u) -> let mib,oib = Global.lookup_inductive ind in (match oib.mind_arity with @@ -49,10 +50,11 @@ let refresh_level evd s = let refresh_universes ?(status=univ_rigid) ?(onlyalg=false) ?(refreshset=false) pbty env evd t = + let open EConstr in let evdref = ref evd in let modified = ref false in let rec refresh status dir t = - match kind_of_term t with + match EConstr.kind !evdref t with | Sort (Type u as s) when (match Univ.universe_level u with | None -> true @@ -72,20 +74,20 @@ let refresh_universes ?(status=univ_rigid) ?(onlyalg=false) ?(refreshset=false) | _ -> t (** Refresh the types of evars under template polymorphic references *) and refresh_term_evars onevars top t = - match kind_of_term (whd_evar !evdref t) with - | App (f, args) when is_template_polymorphic env !evdref (EConstr.of_constr f) -> - let pos = get_polymorphic_positions f in + match EConstr.kind !evdref t with + | App (f, args) when is_template_polymorphic env !evdref f -> + let pos = get_polymorphic_positions !evdref f in refresh_polymorphic_positions args pos - | App (f, args) when top && isEvar f -> + | App (f, args) when top && isEvar !evdref f -> refresh_term_evars true false f; Array.iter (refresh_term_evars onevars false) args | Evar (ev, a) when onevars -> let evi = Evd.find !evdref ev in - let ty' = refresh univ_flexible true evi.evar_concl in + let ty' = refresh univ_flexible true (EConstr.of_constr evi.evar_concl) in if !modified then - evdref := Evd.add !evdref ev {evi with evar_concl = ty'} + evdref := Evd.add !evdref ev {evi with evar_concl = EConstr.Unsafe.to_constr ty'} else () - | _ -> Constr.iter (refresh_term_evars onevars false) t + | _ -> EConstr.iter !evdref (refresh_term_evars onevars false) t and refresh_polymorphic_positions args pos = let rec aux i = function | Some l :: ls -> @@ -100,17 +102,17 @@ let refresh_universes ?(status=univ_rigid) ?(onlyalg=false) ?(refreshset=false) in aux 0 pos in let t' = - if isArity t then + if isArity !evdref t then (match pbty with | None -> t | Some dir -> refresh status dir t) else (refresh_term_evars false true t; t) in - if !modified then !evdref, t' else !evdref, t + if !modified then !evdref, EConstr.Unsafe.to_constr t' else !evdref, EConstr.Unsafe.to_constr t let get_type_of_refresh ?(polyprop=true) ?(lax=false) env sigma c = let ty = Retyping.get_type_of ~polyprop ~lax env sigma c in - refresh_universes (Some false) env sigma ty + refresh_universes (Some false) env sigma (EConstr.of_constr ty) (************************) @@ -127,6 +129,8 @@ let test_success conv_algo env evd c c' rhs = is_success (conv_algo env evd c c' rhs) let add_conv_oriented_pb ?(tail=true) (pbty,env,t1,t2) evd = + let t1 = EConstr.Unsafe.to_constr t1 in + let t2 = EConstr.Unsafe.to_constr t2 in match pbty with | Some true -> add_conv_pb ~tail (Reduction.CUMUL,env,t1,t2) evd | Some false -> add_conv_pb ~tail (Reduction.CUMUL,env,t2,t1) evd @@ -134,29 +138,30 @@ let add_conv_oriented_pb ?(tail=true) (pbty,env,t1,t2) evd = (* We retype applications to ensure the universe constraints are collected *) -exception IllTypedInstance of env * types * types +exception IllTypedInstance of env * EConstr.types * EConstr.types let recheck_applications conv_algo env evdref t = + let open EConstr in let rec aux env t = - match kind_of_term t with + match EConstr.kind !evdref t with | App (f, args) -> let () = aux env f in - let fty = Retyping.get_type_of env !evdref (EConstr.of_constr f) in - let argsty = Array.map (fun x -> aux env x; Retyping.get_type_of env !evdref (EConstr.of_constr x)) args in + let fty = Retyping.get_type_of env !evdref f in + let argsty = Array.map (fun x -> aux env x; Retyping.get_type_of env !evdref x) args in let rec aux i ty = if i < Array.length argsty then - match kind_of_term (whd_all env !evdref (EConstr.of_constr ty)) with + match EConstr.kind !evdref (EConstr.of_constr (whd_all env !evdref ty)) with | Prod (na, dom, codom) -> - (match conv_algo env !evdref Reduction.CUMUL argsty.(i) dom with + (match conv_algo env !evdref Reduction.CUMUL (EConstr.of_constr argsty.(i)) dom with | Success evd -> evdref := evd; - aux (succ i) (subst1 args.(i) codom) + aux (succ i) (Vars.subst1 args.(i) codom) | UnifFailure (evd, reason) -> - Pretype_errors.error_cannot_unify env evd ~reason (argsty.(i), dom)) - | _ -> raise (IllTypedInstance (env, ty, argsty.(i))) + Pretype_errors.error_cannot_unify env evd ~reason (argsty.(i), EConstr.Unsafe.to_constr dom)) + | _ -> raise (IllTypedInstance (env, ty, EConstr.of_constr argsty.(i))) else () - in aux 0 fty + in aux 0 (EConstr.of_constr fty) | _ -> - iter_constr_with_full_binders (fun d env -> push_rel d env) aux env t + iter_with_full_binders !evdref (fun d env -> push_rel d env) aux env t in aux env t @@ -169,7 +174,7 @@ type 'a update = | NoUpdate open Context.Named.Declaration -let inst_of_vars sign = Array.map_of_list (get_id %> mkVar) sign +let inst_of_vars sign = Array.map_of_list (get_id %> EConstr.mkVar) sign let restrict_evar_key evd evk filter candidates = match filter, candidates with @@ -186,7 +191,7 @@ let restrict_evar_key evd evk filter candidates = | Some filter -> filter in let candidates = match candidates with | NoUpdate -> evi.evar_candidates - | UpdateWith c -> Some c in + | UpdateWith c -> Some (List.map EConstr.Unsafe.to_constr c) in let sigma = Sigma.Unsafe.of_evar_map evd in let Sigma (evk, sigma, _) = restrict_evar sigma evk filter candidates in (Sigma.to_evar_map sigma, evk) @@ -216,27 +221,25 @@ let restrict_instance evd evk filter argsv = open Context.Rel.Declaration let noccur_evar env evd evk c = + let open EConstr in let cache = ref Int.Set.empty (* cache for let-ins *) in let rec occur_rec check_types (k, env as acc) c = - match kind_of_term c with + match EConstr.kind evd c with | Evar (evk',args' as ev') -> - (match safe_evar_value evd ev' with - | Some c -> occur_rec check_types acc c - | None -> - if Evar.equal evk evk' then raise Occur - else (if check_types then - occur_rec false acc (existential_type evd ev'); - Array.iter (occur_rec check_types acc) args')) + if Evar.equal evk evk' then raise Occur + else (if check_types then + occur_rec false acc (existential_type evd ev'); + Array.iter (occur_rec check_types acc) args') | Rel i when i > k -> if not (Int.Set.mem (i-k) !cache) then let decl = Environ.lookup_rel i env in if check_types then - (cache := Int.Set.add (i-k) !cache; occur_rec false acc (lift i (get_type decl))); + (cache := Int.Set.add (i-k) !cache; occur_rec false acc (Vars.lift i (EConstr.of_constr (get_type decl)))); (match decl with | LocalAssum _ -> () - | LocalDef (_,b,_) -> cache := Int.Set.add (i-k) !cache; occur_rec false acc (lift i b)) + | LocalDef (_,b,_) -> cache := Int.Set.add (i-k) !cache; occur_rec false acc (Vars.lift i (EConstr.of_constr b))) | Proj (p,c) -> occur_rec true acc c - | _ -> iter_constr_with_full_binders (fun rd (k,env) -> (succ k, push_rel rd env)) + | _ -> iter_with_full_binders evd (fun rd (k,env) -> (succ k, push_rel rd env)) (occur_rec check_types) acc c in try occur_rec false (0,env) c; true with Occur -> false @@ -249,13 +252,14 @@ let noccur_evar env evd evk c = dependencies in variables are canonically associated to the most ancient variable in its family of aliased variables *) -let compute_var_aliases sign = +let compute_var_aliases sign sigma = let open Context.Named.Declaration in List.fold_right (fun decl aliases -> let id = get_id decl in match decl with | LocalDef (_,t,_) -> - (match kind_of_term t with + let t = EConstr.of_constr t in + (match EConstr.kind sigma t with | Var id' -> let aliases_of_id = try Id.Map.find id' aliases with Not_found -> [] in @@ -265,13 +269,16 @@ let compute_var_aliases sign = | LocalAssum _ -> aliases) sign Id.Map.empty -let compute_rel_aliases var_aliases rels = +let compute_rel_aliases var_aliases rels sigma = + let open EConstr in snd (List.fold_right (fun decl (n,aliases) -> (n-1, match decl with | LocalDef (_,t,u) -> - (match kind_of_term t with + let t = EConstr.of_constr t in + let u = EConstr.of_constr u in + (match EConstr.kind sigma t with | Var id' -> let aliases_of_n = try Id.Map.find id' var_aliases with Not_found -> [] in @@ -281,52 +288,57 @@ let compute_rel_aliases var_aliases rels = try Int.Map.find (p+n) aliases with Not_found -> [] in Int.Map.add n (aliases_of_n@[mkRel (p+n)]) aliases | _ -> - Int.Map.add n [lift n (mkCast(t,DEFAULTcast,u))] aliases) + Int.Map.add n [Vars.lift n (mkCast(t,DEFAULTcast,u))] aliases) | LocalAssum _ -> aliases) ) rels (List.length rels,Int.Map.empty)) -let make_alias_map env = +let make_alias_map env sigma = (* We compute the chain of aliases for each var and rel *) - let var_aliases = compute_var_aliases (named_context env) in - let rel_aliases = compute_rel_aliases var_aliases (rel_context env) in + let var_aliases = compute_var_aliases (named_context env) sigma in + let rel_aliases = compute_rel_aliases var_aliases (rel_context env) sigma in (var_aliases,rel_aliases) let lift_aliases n (var_aliases,rel_aliases as aliases) = + let open EConstr in if Int.equal n 0 then aliases else (var_aliases, - Int.Map.fold (fun p l -> Int.Map.add (p+n) (List.map (lift n) l)) + Int.Map.fold (fun p l -> Int.Map.add (p+n) (List.map (Vars.lift n) l)) rel_aliases Int.Map.empty) -let get_alias_chain_of aliases x = match kind_of_term x with +let get_alias_chain_of sigma aliases x = match EConstr.kind sigma x with | Rel n -> (try Int.Map.find n (snd aliases) with Not_found -> []) | Var id -> (try Id.Map.find id (fst aliases) with Not_found -> []) | _ -> [] -let normalize_alias_opt aliases x = - match get_alias_chain_of aliases x with +let normalize_alias_opt sigma aliases x = + let open EConstr in + match get_alias_chain_of sigma aliases x with | [] -> None - | a::_ when isRel a || isVar a -> Some a + | a::_ when isRel sigma a || isVar sigma a -> Some a | [_] -> None | _::a::_ -> Some a -let normalize_alias aliases x = - match normalize_alias_opt aliases x with +let normalize_alias sigma aliases x = + match normalize_alias_opt sigma aliases x with | Some a -> a | None -> x -let normalize_alias_var var_aliases id = - destVar (normalize_alias (var_aliases,Int.Map.empty) (mkVar id)) +let normalize_alias_var sigma var_aliases id = + let open EConstr in + destVar sigma (normalize_alias sigma (var_aliases,Int.Map.empty) (mkVar id)) -let extend_alias decl (var_aliases,rel_aliases) = +let extend_alias sigma decl (var_aliases,rel_aliases) = + let open EConstr in let rel_aliases = - Int.Map.fold (fun n l -> Int.Map.add (n+1) (List.map (lift 1) l)) + Int.Map.fold (fun n l -> Int.Map.add (n+1) (List.map (Vars.lift 1) l)) rel_aliases Int.Map.empty in let rel_aliases = match decl with | LocalDef(_,t,_) -> - (match kind_of_term t with + let t = EConstr.of_constr t in + (match EConstr.kind sigma t with | Var id' -> let aliases_of_binder = try Id.Map.find id' var_aliases with Not_found -> [] in @@ -336,37 +348,38 @@ let extend_alias decl (var_aliases,rel_aliases) = try Int.Map.find (p+1) rel_aliases with Not_found -> [] in Int.Map.add 1 (aliases_of_binder@[mkRel (p+1)]) rel_aliases | _ -> - Int.Map.add 1 [lift 1 t] rel_aliases) + Int.Map.add 1 [Vars.lift 1 t] rel_aliases) | LocalAssum _ -> rel_aliases in (var_aliases, rel_aliases) -let expand_alias_once aliases x = - match get_alias_chain_of aliases x with +let expand_alias_once sigma aliases x = + match get_alias_chain_of sigma aliases x with | [] -> None | l -> Some (List.last l) -let expansions_of_var aliases x = - match get_alias_chain_of aliases x with +let expansions_of_var sigma aliases x = + let open EConstr in + match get_alias_chain_of sigma aliases x with | [] -> [x] - | a::_ as l when isRel a || isVar a -> x :: List.rev l + | a::_ as l when isRel sigma a || isVar sigma a -> x :: List.rev l | _::l -> x :: List.rev l -let expansion_of_var aliases x = - match get_alias_chain_of aliases x with +let expansion_of_var sigma aliases x = + match get_alias_chain_of sigma aliases x with | [] -> x | a::_ -> a -let rec expand_vars_in_term_using sigma aliases t = match kind_of_term t with +let rec expand_vars_in_term_using sigma aliases t = match EConstr.kind sigma t with | Rel _ | Var _ -> - normalize_alias aliases t + normalize_alias sigma aliases t | _ -> - let self aliases c = EConstr.of_constr (expand_vars_in_term_using sigma aliases (EConstr.Unsafe.to_constr c)) in - EConstr.Unsafe.to_constr (map_constr_with_full_binders sigma - extend_alias self aliases (EConstr.of_constr t)) + let self aliases c = expand_vars_in_term_using sigma aliases c in + map_constr_with_full_binders sigma (extend_alias sigma) self aliases t -let expand_vars_in_term env sigma = expand_vars_in_term_using sigma (make_alias_map env) +let expand_vars_in_term env sigma = expand_vars_in_term_using sigma (make_alias_map env sigma) -let free_vars_and_rels_up_alias_expansion aliases c = +let free_vars_and_rels_up_alias_expansion sigma aliases c = + let open EConstr in let acc1 = ref Int.Set.empty and acc2 = ref Id.Set.empty in let acc3 = ref Int.Set.empty and acc4 = ref Id.Set.empty in let cache_rel = ref Int.Set.empty and cache_var = ref Id.Set.empty in @@ -379,25 +392,25 @@ let free_vars_and_rels_up_alias_expansion aliases c = | Var s -> cache_var := Id.Set.add s !cache_var | _ -> () in let rec frec (aliases,depth) c = - match kind_of_term c with + match EConstr.kind sigma c with | Rel _ | Var _ as ck -> if is_in_cache depth ck then () else begin put_in_cache depth ck; - let c' = expansion_of_var aliases c in - (if c != c' then (* expansion, hence a let-in *) - match kind_of_term c with + let c' = expansion_of_var sigma aliases c in + (if c != c' then (* expansion, hence a let-in *) (** FIXME *) + match EConstr.kind sigma c with | Var id -> acc4 := Id.Set.add id !acc4 | Rel n -> if n >= depth+1 then acc3 := Int.Set.add (n-depth) !acc3 | _ -> ()); - match kind_of_term c' with + match EConstr.kind sigma c' with | Var id -> acc2 := Id.Set.add id !acc2 | Rel n -> if n >= depth+1 then acc1 := Int.Set.add (n-depth) !acc1 | _ -> frec (aliases,depth) c end | Const _ | Ind _ | Construct _ -> - acc2 := Id.Set.union (vars_of_global (Global.env()) c) !acc2 + acc2 := Id.Set.union (vars_of_global (Global.env()) (EConstr.to_constr sigma c)) !acc2 | _ -> - iter_constr_with_full_binders - (fun d (aliases,depth) -> (extend_alias d aliases,depth+1)) + iter_with_full_binders sigma + (fun d (aliases,depth) -> (extend_alias sigma d aliases,depth+1)) frec (aliases,depth) c in frec (aliases,0) c; @@ -407,11 +420,11 @@ let free_vars_and_rels_up_alias_expansion aliases c = (* Managing pattern-unification *) (********************************) -let rec expand_and_check_vars aliases = function +let rec expand_and_check_vars sigma aliases = function | [] -> [] - | a::l when isRel a || isVar a -> - let a = expansion_of_var aliases a in - if isRel a || isVar a then a :: expand_and_check_vars aliases l + | a::l when EConstr.isRel sigma a || EConstr.isVar sigma a -> + let a = expansion_of_var sigma aliases a in + if EConstr.isRel sigma a || EConstr.isVar sigma a then a :: expand_and_check_vars sigma aliases l else raise Exit | _ -> raise Exit @@ -422,24 +435,25 @@ module Constrhash = Hashtbl.Make let hash = hash_constr end) -let constr_list_distinct l = +let constr_list_distinct sigma l = let visited = Constrhash.create 23 in let rec loop = function | h::t -> + let h = EConstr.to_constr sigma h in if Constrhash.mem visited h then false else (Constrhash.add visited h h; loop t) | [] -> true in loop l let get_actual_deps evd aliases l t = - if occur_meta_or_existential evd (EConstr.of_constr t) then + if occur_meta_or_existential evd t then (* Probably no restrictions on allowed vars in presence of evars *) l else (* Probably strong restrictions coming from t being evar-closed *) - let (fv_rels,fv_ids,_,_) = free_vars_and_rels_up_alias_expansion aliases t in + let (fv_rels,fv_ids,_,_) = free_vars_and_rels_up_alias_expansion evd aliases t in List.filter (fun c -> - match kind_of_term c with + match EConstr.kind evd c with | Var id -> Id.Set.mem id fv_ids | Rel n -> Int.Set.mem n fv_rels | _ -> assert false) l @@ -462,10 +476,11 @@ let remove_instance_local_defs evd evk args = (* Check if an applied evar "?X[args] l" is a Miller's pattern *) let find_unification_pattern_args env evd l t = - if List.for_all (fun x -> isRel x || isVar x) l (* common failure case *) then - let aliases = make_alias_map env in - match (try Some (expand_and_check_vars aliases l) with Exit -> None) with - | Some l as x when constr_list_distinct (get_actual_deps evd aliases l t) -> x + let open EConstr in + if List.for_all (fun x -> isRel evd x || isVar evd x) l (* common failure case *) then + let aliases = make_alias_map env evd in + match (try Some (expand_and_check_vars evd aliases l) with Exit -> None) with + | Some l as x when constr_list_distinct evd (get_actual_deps evd aliases l t) -> x | _ -> None else None @@ -473,15 +488,17 @@ let find_unification_pattern_args env evd l t = let is_unification_pattern_meta env evd nb m l t = (* Variables from context and rels > nb are implicitly all there *) (* so we need to be a rel <= nb *) - if List.for_all (fun x -> isRel x && destRel x <= nb) l then + let open EConstr in + if List.for_all (fun x -> isRel evd x && destRel evd x <= nb) l then match find_unification_pattern_args env evd l t with - | Some _ as x when not (dependent evd (EConstr.mkMeta m) (EConstr.of_constr t)) -> x + | Some _ as x when not (dependent evd (EConstr.mkMeta m) t) -> x | _ -> None else None let is_unification_pattern_evar env evd (evk,args) l t = - if List.for_all (fun x -> isRel x || isVar x) l + let open EConstr in + if List.for_all (fun x -> isRel evd x || isVar evd x) l && noccur_evar env evd evk t then let args = remove_instance_local_defs evd evk args in @@ -498,7 +515,7 @@ let is_unification_pattern_pure_evar env evd (evk,args) t = | Some _ -> true let is_unification_pattern (env,nb) evd f l t = - match kind_of_term f with + match EConstr.kind evd f with | Meta m -> is_unification_pattern_meta env evd nb m l t | Evar ev -> is_unification_pattern_evar env evd ev l t | _ -> None @@ -511,21 +528,23 @@ let is_unification_pattern (env,nb) evd f l t = dependency: ?X x = ?1 (?1 is a meta) will return \_.?1 while it should return \y. ?1{x\y} (non constant function if ?1 depends on x) (BB) *) let solve_pattern_eqn env sigma l c = + let open EConstr in let c' = List.fold_right (fun a c -> - let c' = subst_term sigma (EConstr.of_constr (lift 1 a)) (EConstr.of_constr (lift 1 c)) in - match kind_of_term a with + let c' = subst_term sigma (Vars.lift 1 a) (Vars.lift 1 c) in + match EConstr.kind sigma a with (* Rem: if [a] links to a let-in, do as if it were an assumption *) | Rel n -> let open Context.Rel.Declaration in let d = map_constr (lift n) (lookup_rel n env) in + let c' = EConstr.of_constr c' in mkLambda_or_LetIn d c' | Var id -> - let d = lookup_named id env in mkNamedLambda_or_LetIn d c' + let d = lookup_named id env in EConstr.of_constr (mkNamedLambda_or_LetIn d c') | _ -> assert false) l c in (* Warning: we may miss some opportunity to eta-reduce more since c' is not in normal form *) - shrink_eta (EConstr.of_constr c') + shrink_eta c' (*****************************************) (* Refining/solving unification problems *) @@ -543,35 +562,33 @@ let solve_pattern_eqn env sigma l c = let make_projectable_subst aliases sigma evi args = let sign = evar_filtered_context evi in - let evar_aliases = compute_var_aliases sign in + let evar_aliases = compute_var_aliases sign sigma in let (_,full_subst,cstr_subst) = List.fold_right (fun decl (args,all,cstrs) -> match decl,args with | LocalAssum (id,c), a::rest -> - let a = whd_evar sigma a in let cstrs = - let a',args = decompose_app_vect sigma (EConstr.of_constr a) in - match kind_of_term a' with + let a',args = decompose_app_vect sigma a in + match EConstr.kind sigma (EConstr.of_constr a') with | Construct cstr -> let l = try Constrmap.find (fst cstr) cstrs with Not_found -> [] in Constrmap.add (fst cstr) ((args,id)::l) cstrs | _ -> cstrs in - (rest,Id.Map.add id [a,normalize_alias_opt aliases a,id] all,cstrs) + (rest,Id.Map.add id [a,normalize_alias_opt sigma aliases a,id] all,cstrs) | LocalDef (id,c,_), a::rest -> - let a = whd_evar sigma a in - (match kind_of_term c with + (match EConstr.kind sigma (EConstr.of_constr c) with | Var id' -> - let idc = normalize_alias_var evar_aliases id' in + let idc = normalize_alias_var sigma evar_aliases id' in let sub = try Id.Map.find idc all with Not_found -> [] in - if List.exists (fun (c,_,_) -> Term.eq_constr a c) sub then + if List.exists (fun (c,_,_) -> EConstr.eq_constr sigma a c) sub then (rest,all,cstrs) else (rest, - Id.Map.add idc ((a,normalize_alias_opt aliases a,id)::sub) all, + Id.Map.add idc ((a,normalize_alias_opt sigma aliases a,id)::sub) all, cstrs) | _ -> - (rest,Id.Map.add id [a,normalize_alias_opt aliases a,id] all,cstrs)) + (rest,Id.Map.add id [a,normalize_alias_opt sigma aliases a,id] all,cstrs)) | _ -> anomaly (Pp.str "Instance does not match its signature")) sign (Array.rev_to_list args,Id.Map.empty,Constrmap.empty) in (full_subst,cstr_subst) @@ -587,15 +604,18 @@ let make_projectable_subst aliases sigma evi args = *) let define_evar_from_virtual_equation define_fun env evd src t_in_env ty_t_in_sign sign filter inst_in_env = + let open EConstr in let evd = Sigma.Unsafe.of_evar_map evd in - let Sigma (evar_in_env, evd, _) = new_evar_instance sign evd ty_t_in_sign ~filter ~src inst_in_env in + let Sigma (evar_in_env, evd, _) = new_evar_instance sign evd ty_t_in_sign ~filter ~src (List.map EConstr.Unsafe.to_constr inst_in_env) in let evd = Sigma.to_evar_map evd in - let t_in_env = whd_evar evd t_in_env in - let evd = define_fun env evd None (destEvar evar_in_env) t_in_env in + let t_in_env = EConstr.of_constr (whd_evar evd (EConstr.Unsafe.to_constr t_in_env)) in + let evar_in_env = EConstr.of_constr evar_in_env in + let (evk, _) = destEvar evd evar_in_env in + let evd = define_fun env evd None (EConstr.destEvar evd evar_in_env) t_in_env in let ctxt = named_context_of_val sign in let inst_in_sign = inst_of_vars (Filter.filter_list filter ctxt) in - let evar_in_sign = mkEvar (fst (destEvar evar_in_env), inst_in_sign) in - (evd,whd_evar evd evar_in_sign) + let evar_in_sign = mkEvar (evk, inst_in_sign) in + (evd,whd_evar evd (EConstr.Unsafe.to_constr evar_in_sign)) (* We have x1..xq |- ?e1 : τ and had to solve something like * Σ; Γ |- ?e1[u1..uq] = (...\y1 ... \yk ... c), where c is typically some @@ -617,10 +637,11 @@ let define_evar_from_virtual_equation define_fun env evd src t_in_env ty_t_in_si exception MorePreciseOccurCheckNeeeded let materialize_evar define_fun env evd k (evk1,args1) ty_in_env = + let open EConstr in if Evd.is_defined evd evk1 then (* Some circularity somewhere (see e.g. #3209) *) raise MorePreciseOccurCheckNeeeded; - let (evk1,args1) = destEvar (whd_evar evd (mkEvar (evk1,args1))) in + let (evk1,args1) = EConstr.destEvar evd (EConstr.mkEvar (evk1,args1)) in let evi1 = Evd.find_undefined evd evk1 in let env1,rel_sign = env_rel_context_chop k env in let sign1 = evar_hyps evi1 in @@ -634,36 +655,38 @@ let materialize_evar define_fun env evd k (evk1,args1) ty_in_env = let LocalAssum (na,t_in_env) | LocalDef (na,_,t_in_env) = d in let id = next_name_away na avoid in let evd,t_in_sign = - let s = Retyping.get_sort_of env evd (EConstr.of_constr t_in_env) in + let t_in_env = EConstr.of_constr t_in_env in + let s = Retyping.get_sort_of env evd t_in_env in let evd,ty_t_in_sign = refresh_universes - ~status:univ_flexible (Some false) env evd (mkSort s) in + ~status:univ_flexible (Some false) env evd (EConstr.mkSort s) in define_evar_from_virtual_equation define_fun env evd src t_in_env ty_t_in_sign sign filter inst_in_env in let evd,d' = match d with | LocalAssum _ -> evd, Context.Named.Declaration.LocalAssum (id,t_in_sign) | LocalDef (_,b,_) -> + let b = EConstr.of_constr b in let evd,b = define_evar_from_virtual_equation define_fun env evd src b t_in_sign sign filter inst_in_env in evd, Context.Named.Declaration.LocalDef (id,b,t_in_sign) in (push_named_context_val d' sign, Filter.extend 1 filter, - (mkRel 1)::(List.map (lift 1) inst_in_env), - (mkRel 1)::(List.map (lift 1) inst_in_sign), + (mkRel 1)::(List.map (Vars.lift 1) inst_in_env), + (mkRel 1)::(List.map (Vars.lift 1) inst_in_sign), push_rel d env,evd,id::avoid)) rel_sign (sign1,filter1,Array.to_list args1,inst_in_sign,env1,evd,ids1) in let evd,ev2ty_in_sign = - let s = Retyping.get_sort_of env evd (EConstr.of_constr ty_in_env) in + let s = Retyping.get_sort_of env evd ty_in_env in let evd,ty_t_in_sign = refresh_universes - ~status:univ_flexible (Some false) env evd (mkSort s) in + ~status:univ_flexible (Some false) env evd (EConstr.mkSort s) in define_evar_from_virtual_equation define_fun env evd src ty_in_env ty_t_in_sign sign2 filter2 inst2_in_env in let evd = Sigma.Unsafe.of_evar_map evd in let Sigma (ev2_in_sign, evd, _) = - new_evar_instance sign2 evd ev2ty_in_sign ~filter:filter2 ~src inst2_in_sign in + new_evar_instance sign2 evd ev2ty_in_sign ~filter:filter2 ~src (List.map EConstr.Unsafe.to_constr inst2_in_sign) in let evd = Sigma.to_evar_map evd in - let ev2_in_env = (fst (destEvar ev2_in_sign), Array.of_list inst2_in_env) in - (evd, ev2_in_sign, ev2_in_env) + let ev2_in_env = (fst (destEvar evd (EConstr.of_constr ev2_in_sign)), Array.of_list inst2_in_env) in + (evd, EConstr.of_constr ev2_in_sign, ev2_in_env) let restrict_upon_filter evd evk p args = let oldfullfilter = evar_filter (Evd.find_undefined evd evk) in @@ -722,7 +745,7 @@ let find_projectable_constructor env evd cstr k args cstr_subst = type evar_projection = | ProjectVar -| ProjectEvar of existential * evar_info * Id.t * evar_projection +| ProjectEvar of EConstr.existential * evar_info * Id.t * evar_projection exception NotUnique exception NotUniqueInType of (Id.t * evar_projection) list @@ -730,19 +753,19 @@ exception NotUniqueInType of (Id.t * evar_projection) list let rec assoc_up_to_alias sigma aliases y yc = function | [] -> raise Not_found | (c,cc,id)::l -> - let c' = whd_evar sigma c in - if Term.eq_constr y c' then id + if EConstr.eq_constr sigma y c then id else match l with | _ :: _ -> assoc_up_to_alias sigma aliases y yc l | [] -> (* Last chance, we reason up to alias conversion *) - match (if c == c' then cc else normalize_alias_opt aliases c') with - | Some cc when Term.eq_constr yc cc -> id - | _ -> if Term.eq_constr yc c then id else raise Not_found + match (normalize_alias_opt sigma aliases c) with + | Some cc when EConstr.eq_constr sigma yc cc -> id + | _ -> if EConstr.eq_constr sigma yc c then id else raise Not_found let rec find_projectable_vars with_evars aliases sigma y subst = - let yc = normalize_alias aliases y in + let open EConstr in + let yc = normalize_alias sigma aliases y in let is_projectable idc idcl subst' = (* First test if some [id] aliased to [idc] is bound to [y] in [subst] *) try @@ -752,12 +775,12 @@ let rec find_projectable_vars with_evars aliases sigma y subst = (* Then test if [idc] is (indirectly) bound in [subst] to some evar *) (* projectable on [y] *) if with_evars then - let f (c,_,id) = isEvar c && is_undefined sigma (fst (destEvar c)) in + let f (c,_,id) = isEvar sigma c in let idcl' = List.filter f idcl in match idcl' with | [c,_,id] -> begin - let (evk,argsv as t) = destEvar c in + let (evk,argsv as t) = destEvar sigma c in let evi = Evd.find sigma evk in let subst,_ = make_projectable_subst aliases sigma evi argsv in let l = find_projectable_vars with_evars aliases sigma y subst in @@ -805,19 +828,19 @@ let rec find_solution_type evarenv = function let rec do_projection_effects define_fun env ty evd = function | ProjectVar -> evd | ProjectEvar ((evk,argsv),evi,id,p) -> - let evd = Evd.define evk (mkVar id) evd in + let open EConstr in + let evd = Evd.define evk (Constr.mkVar id) evd in (* TODO: simplify constraints involving evk *) let evd = do_projection_effects define_fun env ty evd p in - let ty = whd_all env evd (EConstr.of_constr (Lazy.force ty)) in - if not (isSort ty) then + let ty = EConstr.of_constr (whd_all env evd (Lazy.force ty)) in + if not (isSort evd ty) then (* Don't try to instantiate if a sort because if evar_concl is an evar it may commit to a univ level which is not the right one (however, regarding coercions, because t is obtained by unif, we know that no coercion can be inserted) *) let subst = make_pure_subst evi argsv in - let ty' = replace_vars subst evi.evar_concl in - let ty' = whd_evar evd ty' in - if isEvar ty' then define_fun env evd (Some false) (destEvar ty') ty else evd + let ty' = Vars.replace_vars subst (EConstr.of_constr evi.evar_concl) in + if isEvar evd ty' then define_fun env evd (Some false) (destEvar evd ty') ty else evd else evd @@ -843,7 +866,7 @@ let rec do_projection_effects define_fun env ty evd = function type projectibility_kind = | NoUniqueProjection - | UniqueProjection of constr * evar_projection list + | UniqueProjection of EConstr.constr * evar_projection list type projectibility_status = | CannotInvert @@ -851,16 +874,16 @@ type projectibility_status = let invert_arg_from_subst evd aliases k0 subst_in_env_extended_with_k_binders c_in_env_extended_with_k_binders = let effects = ref [] in + let open EConstr in let rec aux k t = - let t = whd_evar evd t in - match kind_of_term t with + match EConstr.kind evd t with | Rel i when i>k0+k -> aux' k (mkRel (i-k)) | Var id -> aux' k t - | _ -> map_constr_with_binders succ aux k t + | _ -> map_with_binders evd succ aux k t and aux' k t = - try project_with_effects aliases evd effects t subst_in_env_extended_with_k_binders + try EConstr.of_constr (project_with_effects aliases evd effects t subst_in_env_extended_with_k_binders) with Not_found -> - match expand_alias_once aliases t with + match expand_alias_once evd aliases t with | None -> raise Not_found | Some c -> aux k c in try @@ -895,7 +918,7 @@ let extract_unique_projection = function let extract_candidates sols = try UpdateWith - (List.map (function (id,ProjectVar) -> mkVar id | _ -> raise Exit) sols) + (List.map (function (id,ProjectVar) -> EConstr.mkVar id | _ -> raise Exit) sols) with Exit -> NoUpdate @@ -929,12 +952,12 @@ let filter_effective_candidates evd evi filter candidates = | None -> candidates | Some filter -> let ids = set_of_evctx (Filter.filter_list filter (evar_context evi)) in - List.filter (fun a -> Id.Set.subset (collect_vars evd (EConstr.of_constr a)) ids) candidates + List.filter (fun a -> Id.Set.subset (collect_vars evd a) ids) candidates let filter_candidates evd evk filter candidates_update = let evi = Evd.find_undefined evd evk in let candidates = match candidates_update with - | NoUpdate -> evi.evar_candidates + | NoUpdate -> Option.map (fun l -> List.map EConstr.of_constr l) evi.evar_candidates | UpdateWith c -> Some c in match candidates with @@ -982,17 +1005,18 @@ let restrict_hyps evd evk filter candidates = let typablefilter = closure_of_filter evd evk (Some filter) in (typablefilter,candidates) -exception EvarSolvedWhileRestricting of evar_map * constr +exception EvarSolvedWhileRestricting of evar_map * EConstr.constr let do_restrict_hyps evd (evk,args as ev) filter candidates = + let open EConstr in let filter,candidates = match filter with | None -> None,candidates | Some filter -> restrict_hyps evd evk filter candidates in match candidates,filter with | UpdateWith [], _ -> error "Not solvable." | UpdateWith [nc],_ -> - let evd = Evd.define evk nc evd in - raise (EvarSolvedWhileRestricting (evd,whd_evar evd (mkEvar ev))) + let evd = Evd.define evk (EConstr.Unsafe.to_constr nc) evd in + raise (EvarSolvedWhileRestricting (evd,mkEvar ev)) | NoUpdate, None -> evd,ev | _ -> restrict_applied_evar evd ev filter candidates @@ -1000,6 +1024,7 @@ let do_restrict_hyps evd (evk,args as ev) filter candidates = (* ?e is assumed to have no candidates *) let postpone_non_unique_projection env evd pbty (evk,argsv as ev) sols rhs = + let open EConstr in let rhs = expand_vars_in_term env evd rhs in let filter = restrict_upon_filter evd evk @@ -1010,8 +1035,8 @@ let postpone_non_unique_projection env evd pbty (evk,argsv as ev) sols rhs = (* that says that the body is hidden. Note that expand_vars_in_term *) (* expands only rels and vars aliases, not rels or vars bound to an *) (* arbitrary complex term *) - (fun a -> not (isRel a || isVar a) - || dependent evd (EConstr.of_constr a) (EConstr.of_constr rhs) || List.exists (fun (id,_) -> isVarId id a) sols) + (fun a -> not (isRel evd a || isVar evd a) + || dependent evd a rhs || List.exists (fun (id,_) -> isVarId evd id a) sols) argsv in let filter = closure_of_filter evd evk filter in let candidates = extract_candidates sols in @@ -1042,6 +1067,9 @@ let postpone_non_unique_projection env evd pbty (evk,argsv as ev) sols rhs = * Note: argument f is the function used to instantiate evars. *) +let instantiate_evar_array evi c args = + EConstr.of_constr (instantiate_evar_array evi (EConstr.Unsafe.to_constr c) (Array.map EConstr.Unsafe.to_constr args)) + let filter_compatible_candidates conv_algo env evd evi args rhs c = let c' = instantiate_evar_array evi c args in match conv_algo env evd Reduction.CONV rhs c' with @@ -1061,6 +1089,8 @@ let restrict_candidates conv_algo env evd filter1 (evk1,argsv1) (evk2,argsv2) = | _, None -> filter_candidates evd evk1 filter1 NoUpdate | None, Some _ -> raise DoesNotPreserveCandidateRestriction | Some l1, Some l2 -> + let l1 = List.map EConstr.of_constr l1 in + let l2 = List.map EConstr.of_constr l2 in let l1 = filter_effective_candidates evd evi1 filter1 l1 in let l1' = List.filter (fun c1 -> let c1' = instantiate_evar_array evi1 c1 argsv1 in @@ -1075,7 +1105,7 @@ let restrict_candidates conv_algo env evd filter1 (evk1,argsv1) (evk2,argsv2) = if Int.equal (List.length l1) (List.length l1') then NoUpdate else UpdateWith l1' -exception CannotProject of evar_map * existential +exception CannotProject of evar_map * EConstr.existential (* Assume that FV(?n[x1:=t1..xn:=tn]) belongs to some set U. Can ?n be instantiated by a term u depending essentially on xi such that the @@ -1092,15 +1122,15 @@ exception CannotProject of evar_map * existential *) let rec is_constrainable_in top evd k (ev,(fv_rels,fv_ids) as g) t = - let f,args = decompose_app_vect evd (EConstr.of_constr t) in - match kind_of_term f with + let f,args = decompose_app_vect evd t in + match EConstr.kind evd (EConstr.of_constr f) with | Construct ((ind,_),u) -> let n = Inductiveops.inductive_nparams ind in if n > Array.length args then true (* We don't try to be more clever *) else let params = fst (Array.chop n args) in - Array.for_all (is_constrainable_in false evd k g) params - | Ind _ -> Array.for_all (is_constrainable_in false evd k g) args + Array.for_all (EConstr.of_constr %> is_constrainable_in false evd k g) params + | Ind _ -> Array.for_all (EConstr.of_constr %> is_constrainable_in false evd k g) args | Prod (na,t1,t2) -> is_constrainable_in false evd k g t1 && is_constrainable_in false evd k g t2 | Evar (ev',_) -> top || not (Evar.equal ev' ev) (*If ev' needed, one may also try to restrict it*) | Var id -> Id.Set.mem id fv_ids @@ -1109,30 +1139,31 @@ let rec is_constrainable_in top evd k (ev,(fv_rels,fv_ids) as g) t = | _ -> (* We don't try to be more clever *) true let has_constrainable_free_vars env evd aliases force k ev (fv_rels,fv_ids,let_rels,let_ids) t = - let t' = expansion_of_var aliases t in + let t' = expansion_of_var evd aliases t in if t' != t then (* t is a local definition, we keep it only if appears in the list *) (* of let-in variables effectively occurring on the right-hand side, *) (* which is the only reason to keep it when inverting arguments *) - match kind_of_term t with + match EConstr.kind evd t with | Var id -> Id.Set.mem id let_ids | Rel n -> Int.Set.mem n let_rels | _ -> assert false else (* t is an instance for a proper variable; we filter it along *) (* the free variables allowed to occur *) - match kind_of_term t with + match EConstr.kind evd t with | Var id -> Id.Set.mem id fv_ids | Rel n -> n <= k || Int.Set.mem n fv_rels | _ -> (not force || noccur_evar env evd ev t) && is_constrainable_in true evd k (ev,(fv_rels,fv_ids)) t -exception EvarSolvedOnTheFly of evar_map * constr +exception EvarSolvedOnTheFly of evar_map * EConstr.constr (* Try to project evk1[argsv1] on evk2[argsv2], if [ev1] is a pattern on the common domain of definition *) let project_evar_on_evar force g env evd aliases k2 pbty (evk1,argsv1 as ev1) (evk2,argsv2 as ev2) = + let open EConstr in (* Apply filtering on ev1 so that fvs(ev1) are in fvs(ev2). *) - let fvs2 = free_vars_and_rels_up_alias_expansion aliases (mkEvar ev2) in + let fvs2 = free_vars_and_rels_up_alias_expansion evd aliases (mkEvar ev2) in let filter1 = restrict_upon_filter evd evk1 (has_constrainable_free_vars env evd aliases force k2 evk2 fvs2) argsv1 in @@ -1161,12 +1192,12 @@ let check_evar_instance evd evk1 body conv_algo = (* FIXME: The body might be ill-typed when this is called from w_merge *) (* This happens in practice, cf MathClasses build failure on 2013-3-15 *) let ty = - try Retyping.get_type_of ~lax:true evenv evd (EConstr.of_constr body) + try EConstr.of_constr (Retyping.get_type_of ~lax:true evenv evd body) with Retyping.RetypeError _ -> error "Ill-typed evar instance" in - match conv_algo evenv evd Reduction.CUMUL ty evi.evar_concl with + match conv_algo evenv evd Reduction.CUMUL ty (EConstr.of_constr evi.evar_concl) with | Success evd -> evd - | UnifFailure _ -> raise (IllTypedInstance (evenv,ty,evi.evar_concl)) + | UnifFailure _ -> raise (IllTypedInstance (evenv,ty,EConstr.of_constr evi.evar_concl)) let update_evar_source ev1 ev2 evd = let loc, evs2 = evar_source ev2 evd in @@ -1178,9 +1209,10 @@ let update_evar_source ev1 ev2 evd = let solve_evar_evar_l2r force f g env evd aliases pbty ev1 (evk2,_ as ev2) = try + let open EConstr in let evd,body = project_evar_on_evar force g env evd aliases 0 pbty ev1 ev2 in - let evd' = Evd.define evk2 body evd in - let evd' = update_evar_source (fst (destEvar body)) evk2 evd' in + let evd' = Evd.define evk2 (EConstr.Unsafe.to_constr body) evd in + let evd' = update_evar_source (fst (destEvar evd body)) evk2 evd' in check_evar_instance evd' evk2 body g with EvarSolvedOnTheFly (evd,c) -> f env evd pbty ev2 c @@ -1197,7 +1229,8 @@ let preferred_orientation evd evk1 evk2 = | _ -> true let solve_evar_evar_aux force f g env evd pbty (evk1,args1 as ev1) (evk2,args2 as ev2) = - let aliases = make_alias_map env in + let open EConstr in + let aliases = make_alias_map env evd in if preferred_orientation evd evk1 evk2 then try solve_evar_evar_l2r force f g env evd aliases (opp_problem pbty) ev2 ev1 with CannotProject (evd,ev2) -> @@ -1244,10 +1277,10 @@ let solve_evar_evar ?(force=false) f g env evd pbty (evk1,args1 as ev1) (evk2,ar solve_evar_evar_aux force f g env evd pbty ev1 ev2 type conv_fun = - env -> evar_map -> conv_pb -> constr -> constr -> unification_result + env -> evar_map -> conv_pb -> EConstr.constr -> EConstr.constr -> unification_result type conv_fun_bool = - env -> evar_map -> conv_pb -> constr -> constr -> bool + env -> evar_map -> conv_pb -> EConstr.constr -> EConstr.constr -> bool (* Solve pbs ?e[t1..tn] = ?e[u1..un] which arise often in fixpoint * definitions. We try to unify the ti with the ui pairwise. The pairs @@ -1255,8 +1288,9 @@ type conv_fun_bool = * depend on these args). *) let solve_refl ?(can_drop=false) conv_algo env evd pbty evk argsv1 argsv2 = + let open EConstr in let evdref = ref evd in - if Array.equal (e_eq_constr_univs evdref) argsv1 argsv2 then !evdref else + if Array.equal (fun c1 c2 -> e_eq_constr_univs evdref (EConstr.Unsafe.to_constr c1) (EConstr.Unsafe.to_constr c2) ) argsv1 argsv2 then !evdref else (* Filter and restrict if needed *) let args = Array.map2 (fun a1 a2 -> (a1, a2)) argsv1 argsv2 in let untypedfilter = @@ -1288,14 +1322,14 @@ let solve_candidates conv_algo env evd (evk,argsv) rhs = | Some l -> let l' = List.map_filter - (filter_compatible_candidates conv_algo env evd evi argsv rhs) l in + (fun c -> filter_compatible_candidates conv_algo env evd evi argsv rhs (EConstr.of_constr c)) l in match l' with | [] -> raise IncompatibleCandidates | [c,evd] -> (* solve_candidates might have been called recursively in the mean *) (* time and the evar been solved by the filtering process *) if Evd.is_undefined evd evk then - let evd' = Evd.define evk c evd in + let evd' = Evd.define evk (EConstr.Unsafe.to_constr c) evd in check_evar_instance evd' evk c conv_algo else evd | l when List.length l < List.length l' -> @@ -1304,7 +1338,9 @@ let solve_candidates conv_algo env evd (evk,argsv) rhs = | l -> evd let occur_evar_upto_types sigma n c = + let c = EConstr.Unsafe.to_constr c in let seen = ref Evar.Set.empty in + (** FIXME: Is that supposed to be evar-insensitive? *) let rec occur_rec c = match kind_of_term c with | Evar (sp,_) when Evar.equal sp n -> raise Occur | Evar (sp,args as e) -> @@ -1341,14 +1377,15 @@ let occur_evar_upto_types sigma n c = * Σ; Γ, y1..yk |- φ(u1..un) = c /\ x1..xn |- φ(x1..xn) *) -exception NotInvertibleUsingOurAlgorithm of constr +exception NotInvertibleUsingOurAlgorithm of EConstr.constr exception NotEnoughInformationToProgress of (Id.t * evar_projection) list -exception NotEnoughInformationEvarEvar of constr -exception OccurCheckIn of evar_map * constr +exception NotEnoughInformationEvarEvar of EConstr.constr +exception OccurCheckIn of evar_map * EConstr.constr exception MetaOccurInBodyInternal let rec invert_definition conv_algo choose env evd pbty (evk,argsv as ev) rhs = - let aliases = make_alias_map env in + let open EConstr in + let aliases = make_alias_map env evd in let evdref = ref evd in let progress = ref false in let evi = Evd.find evd evk in @@ -1365,7 +1402,7 @@ let rec invert_definition conv_algo choose env evd pbty (evk,argsv as ev) rhs = | (id,p)::_::_ -> if choose then (mkVar id, p) else raise (NotUniqueInType sols) in - let ty = lazy (Retyping.get_type_of env !evdref (EConstr.of_constr t)) in + let ty = lazy (EConstr.of_constr (Retyping.get_type_of env !evdref t)) in let evd = do_projection_effects (evar_define conv_algo ~choose) env ty !evdref p in evdref := evd; c @@ -1377,18 +1414,18 @@ let rec invert_definition conv_algo choose env evd pbty (evk,argsv as ev) rhs = (* No unique projection but still restrict to where it is possible *) (* materializing is necessary, but is restricting useful? *) let ty = find_solution_type (evar_filtered_env evi) sols in - let ty' = instantiate_evar_array evi ty argsv in + let ty' = instantiate_evar_array evi (EConstr.of_constr ty) argsv in let (evd,evar,(evk',argsv' as ev')) = materialize_evar (evar_define conv_algo ~choose) env !evdref 0 ev ty' in - let ts = expansions_of_var aliases t in - let test c = isEvar c || List.mem_f Constr.equal c ts in + let ts = expansions_of_var evd aliases t in + let test c = isEvar evd c || List.mem_f (EConstr.eq_constr evd) c ts in let filter = restrict_upon_filter evd evk test argsv' in let filter = closure_of_filter evd evk' filter in let candidates = extract_candidates sols in let evd = match candidates with | NoUpdate -> let evd, ev'' = restrict_applied_evar evd ev' filter NoUpdate in - Evd.add_conv_pb (Reduction.CONV,env,mkEvar ev'',t) evd + add_conv_oriented_pb ~tail:false (None,env,mkEvar ev'',t) evd | UpdateWith _ -> restrict_evar evd evk' filter candidates in @@ -1396,29 +1433,28 @@ let rec invert_definition conv_algo choose env evd pbty (evk,argsv as ev) rhs = evar in let rec imitate (env',k as envk) t = - let t = whd_evar !evdref t in - match kind_of_term t with + match EConstr.kind !evdref t with | Rel i when i>k -> let open Context.Rel.Declaration in (match Environ.lookup_rel (i-k) env' with | LocalAssum _ -> project_variable (mkRel (i-k)) | LocalDef (_,b,_) -> try project_variable (mkRel (i-k)) - with NotInvertibleUsingOurAlgorithm _ -> imitate envk (lift i b)) + with NotInvertibleUsingOurAlgorithm _ -> imitate envk (Vars.lift i (EConstr.of_constr b))) | Var id -> (match Environ.lookup_named id env' with | LocalAssum _ -> project_variable t | LocalDef (_,b,_) -> try project_variable t - with NotInvertibleUsingOurAlgorithm _ -> imitate envk b) + with NotInvertibleUsingOurAlgorithm _ -> imitate envk (EConstr.of_constr b)) | LetIn (na,b,u,c) -> - imitate envk (subst1 b c) + imitate envk (Vars.subst1 b c) | Evar (evk',args' as ev') -> if Evar.equal evk evk' then raise (OccurCheckIn (evd,rhs)); (* Evar/Evar problem (but left evar is virtual) *) let aliases = lift_aliases k aliases in (try - let ev = (evk,Array.map (lift k) argsv) in + let ev = (evk,Array.map (Vars.lift k) argsv) in let evd,body = project_evar_on_evar false conv_algo env' !evdref aliases k None ev' ev in evdref := evd; body @@ -1428,7 +1464,7 @@ let rec invert_definition conv_algo choose env evd pbty (evk,argsv as ev) rhs = if not !progress then raise (NotEnoughInformationEvarEvar t); (* Make the virtual left evar real *) - let ty = get_type_of env' evd (EConstr.of_constr t) in + let ty = EConstr.of_constr (get_type_of env' evd t) in let (evd,evar'',ev'') = materialize_evar (evar_define conv_algo ~choose) env' evd k ev ty in (* materialize_evar may instantiate ev' by another evar; adjust it *) @@ -1437,7 +1473,7 @@ let rec invert_definition conv_algo choose env evd pbty (evk,argsv as ev) rhs = (* Try to project (a restriction of) the left evar ... *) try let evd,body = project_evar_on_evar false conv_algo env' evd aliases 0 None ev'' ev' in - let evd = Evd.define evk' body evd in + let evd = Evd.define evk' (EConstr.Unsafe.to_constr body) evd in check_evar_instance evd evk' body conv_algo with | EvarSolvedOnTheFly _ -> assert false (* ev has no candidates *) @@ -1449,9 +1485,9 @@ let rec invert_definition conv_algo choose env evd pbty (evk,argsv as ev) rhs = | _ -> progress := true; match - let c,args = decompose_app_vect !evdref (EConstr.of_constr t) in - match kind_of_term c with - | Construct (cstr,u) when noccur_between 1 k t -> + let c,args = decompose_app_vect !evdref t in + match EConstr.kind !evdref (EConstr.of_constr c) with + | Construct (cstr,u) when Vars.noccur_between !evdref 1 k t -> (* This is common case when inferring the return clause of match *) (* (currently rudimentary: we do not treat the case of multiple *) (* possible inversions; we do not treat overlap with a possible *) @@ -1462,14 +1498,13 @@ let rec invert_definition conv_algo choose env evd pbty (evk,argsv as ev) rhs = | _ -> None with | Some l -> - let ty = get_type_of env' !evdref (EConstr.of_constr t) in + let ty = EConstr.of_constr (get_type_of env' !evdref t) in let candidates = try - let self env c = EConstr.of_constr (imitate env (EConstr.Unsafe.to_constr c)) in let t = map_constr_with_full_binders !evdref (fun d (env,k) -> push_rel d env, k+1) - self envk (EConstr.of_constr t) in - EConstr.Unsafe.to_constr t::l + imitate envk t in + t::l with e when CErrors.noncritical e -> l in (match candidates with | [x] -> x @@ -1480,11 +1515,10 @@ let rec invert_definition conv_algo choose env evd pbty (evk,argsv as ev) rhs = evar'') | None -> (* Evar/Rigid problem (or assimilated if not normal): we "imitate" *) - let self env c = EConstr.of_constr (imitate env (EConstr.Unsafe.to_constr c)) in - EConstr.Unsafe.to_constr (map_constr_with_full_binders !evdref (fun d (env,k) -> push_rel d env, k+1) - self envk (EConstr.of_constr t)) + map_constr_with_full_binders !evdref (fun d (env,k) -> push_rel d env, k+1) + imitate envk t in - let rhs = whd_beta evd (EConstr.of_constr rhs) (* heuristic *) in + let rhs = EConstr.of_constr (whd_beta evd rhs) (* heuristic *) in let fast rhs = let filter_ctxt = evar_filtered_context evi in let names = ref Idset.empty in @@ -1493,16 +1527,16 @@ let rec invert_definition conv_algo choose env evd pbty (evk,argsv as ev) rhs = | (decl :: ctxt'), (c :: s') -> let id = get_id decl in names := Idset.add id !names; - isVarId id c && is_id_subst ctxt' s' + isVarId evd id c && is_id_subst ctxt' s' | [], [] -> true | _ -> false in is_id_subst filter_ctxt (Array.to_list argsv) && - closed0 rhs && - Idset.subset (collect_vars evd (EConstr.of_constr rhs)) !names + Vars.closed0 evd rhs && + Idset.subset (collect_vars evd rhs) !names in let body = - if fast rhs then nf_evar evd rhs + if fast rhs then EConstr.of_constr (EConstr.to_constr evd rhs) (** FIXME? *) else let t' = imitate (env,0) rhs in if !progress then @@ -1518,7 +1552,7 @@ let rec invert_definition conv_algo choose env evd pbty (evk,argsv as ev) rhs = *) and evar_define conv_algo ?(choose=false) env evd pbty (evk,argsv as ev) rhs = - match kind_of_term rhs with + match EConstr.kind evd rhs with | Evar (evk2,argsv2 as ev2) -> if Evar.equal evk evk2 then solve_refl ~can_drop:choose @@ -1531,7 +1565,7 @@ and evar_define conv_algo ?(choose=false) env evd pbty (evk,argsv as ev) rhs = with NoCandidates -> try let (evd',body) = invert_definition conv_algo choose env evd pbty ev rhs in - if occur_meta evd' (EConstr.of_constr body) then raise MetaOccurInBodyInternal; + if occur_meta evd' body then raise MetaOccurInBodyInternal; (* invert_definition may have instantiate some evars of rhs with evk *) (* so we recheck acyclicity *) if occur_evar_upto_types evd' evk body then raise (OccurCheckIn (evd',body)); @@ -1553,23 +1587,23 @@ and evar_define conv_algo ?(choose=false) env evd pbty (evk,argsv as ev) rhs = str "----> " ++ int ev ++ str " := " ++ print_constr body); raise e in*) - let evd' = check_evar_instance evd' evk body conv_algo in + let evd' = check_evar_instance evd' evk (EConstr.of_constr body) conv_algo in Evd.define evk body evd' with | NotEnoughInformationToProgress sols -> postpone_non_unique_projection env evd pbty ev sols rhs | NotEnoughInformationEvarEvar t -> - add_conv_oriented_pb (pbty,env,mkEvar ev,t) evd + add_conv_oriented_pb (pbty,env,EConstr.mkEvar ev,t) evd | MorePreciseOccurCheckNeeeded -> - add_conv_oriented_pb (pbty,env,mkEvar ev,rhs) evd + add_conv_oriented_pb (pbty,env,EConstr.mkEvar ev,rhs) evd | NotInvertibleUsingOurAlgorithm _ | MetaOccurInBodyInternal as e -> raise e | OccurCheckIn (evd,rhs) -> (* last chance: rhs actually reduces to ev *) - let c = whd_all env evd (EConstr.of_constr rhs) in - match kind_of_term c with + let c = EConstr.of_constr (whd_all env evd rhs) in + match EConstr.kind evd c with | Evar (evk',argsv2) when Evar.equal evk evk' -> - solve_refl (fun env sigma pb c c' -> is_fconv pb env sigma (EConstr.of_constr c) (EConstr.of_constr c')) + solve_refl (fun env sigma pb c c' -> is_fconv pb env sigma c c') env evd pbty evk argsv argsv2 | _ -> raise (OccurCheckIn (evd,rhs)) @@ -1610,7 +1644,7 @@ let reconsider_conv_pbs conv_algo evd = (fun p (pbty,env,t1,t2 as x) -> match p with | Success evd -> - (match conv_algo env evd pbty t1 t2 with + (match conv_algo env evd pbty (EConstr.of_constr t1) (EConstr.of_constr t2) with | Success _ as x -> x | UnifFailure (i,e) -> UnifFailure (i,CannotSolveConstraint (x,e))) | UnifFailure _ as x -> x) @@ -1624,8 +1658,9 @@ let reconsider_conv_pbs conv_algo evd = (* Rq: uncomplete algorithm if pbty = CONV_X_LEQ ! *) let solve_simple_eqn conv_algo ?(choose=false) env evd (pbty,(evk1,args1 as ev1),t2) = + let open EConstr in try - let t2 = whd_betaiota evd t2 in (* includes whd_evar *) + let t2 = EConstr.of_constr (whd_betaiota evd t2) in (* includes whd_evar *) let evd = evar_define conv_algo ~choose env evd pbty ev1 t2 in reconsider_conv_pbs conv_algo evd with @@ -1638,5 +1673,5 @@ let solve_simple_eqn conv_algo ?(choose=false) env evd (pbty,(evk1,args1 as ev1) | IllTypedInstance (env,t,u) -> UnifFailure (evd,InstanceNotSameType (evk1,env,t,u)) | IncompatibleCandidates -> - UnifFailure (evd,ConversionFailed (env,mkEvar ev1, EConstr.Unsafe.to_constr t2)) + UnifFailure (evd,ConversionFailed (env,mkEvar ev1,t2)) diff --git a/pretyping/evarsolve.mli b/pretyping/evarsolve.mli index ac082d1bf..23cb245e0 100644 --- a/pretyping/evarsolve.mli +++ b/pretyping/evarsolve.mli @@ -7,6 +7,7 @@ (************************************************************************) open Term +open EConstr open Evd open Environ @@ -41,7 +42,7 @@ val refresh_universes : (* Also refresh Prop and Set universes, so that the returned type can be any supertype of the original type *) bool option (* direction: true for levels lower than the existing levels *) -> - env -> evar_map -> types -> evar_map * types + env -> evar_map -> types -> evar_map * Constr.types val solve_refl : ?can_drop:bool -> conv_fun_bool -> env -> evar_map -> bool option -> existential_key -> constr array -> constr array -> evar_map @@ -52,7 +53,7 @@ val solve_evar_evar : ?force:bool -> env -> evar_map -> bool option -> existential -> existential -> evar_map val solve_simple_eqn : conv_fun -> ?choose:bool -> env -> evar_map -> - bool option * existential * EConstr.t -> unification_result + bool option * existential * constr -> unification_result val reconsider_conv_pbs : conv_fun -> evar_map -> unification_result @@ -62,7 +63,7 @@ val is_unification_pattern_evar : env -> evar_map -> existential -> constr list val is_unification_pattern : env * int -> evar_map -> constr -> constr list -> constr -> constr list option -val solve_pattern_eqn : env -> evar_map -> constr list -> constr -> constr +val solve_pattern_eqn : env -> evar_map -> constr list -> constr -> Constr.t val noccur_evar : env -> evar_map -> Evar.t -> constr -> bool @@ -73,7 +74,7 @@ val check_evar_instance : evar_map -> existential_key -> constr -> conv_fun -> evar_map val remove_instance_local_defs : - evar_map -> existential_key -> constr array -> constr list + evar_map -> existential_key -> 'a array -> 'a list val get_type_of_refresh : - ?polyprop:bool -> ?lax:bool -> env -> evar_map -> EConstr.constr -> evar_map * types + ?polyprop:bool -> ?lax:bool -> env -> evar_map -> constr -> evar_map * Constr.types diff --git a/pretyping/pretype_errors.ml b/pretyping/pretype_errors.ml index f28fb84ba..c14d81505 100644 --- a/pretyping/pretype_errors.ml +++ b/pretyping/pretype_errors.ml @@ -13,14 +13,14 @@ open Environ open Type_errors type unification_error = - | OccurCheck of existential_key * constr - | NotClean of existential * env * constr (* Constr is a variable not in scope *) + | OccurCheck of existential_key * EConstr.constr + | NotClean of EConstr.existential * env * EConstr.constr (* Constr is a variable not in scope *) | NotSameArgSize | NotSameHead | NoCanonicalStructure - | ConversionFailed of env * constr * constr (* Non convertible closed terms *) + | ConversionFailed of env * EConstr.constr * EConstr.constr (* Non convertible closed terms *) | MetaOccurInBody of existential_key - | InstanceNotSameType of existential_key * env * types * types + | InstanceNotSameType of existential_key * env * EConstr.types * EConstr.types | UnifUnivInconsistency of Univ.univ_inconsistency | CannotSolveConstraint of Evd.evar_constraint * unification_error | ProblemBeyondCapabilities diff --git a/pretyping/pretype_errors.mli b/pretyping/pretype_errors.mli index 8a6d8b6b3..217deda4d 100644 --- a/pretyping/pretype_errors.mli +++ b/pretyping/pretype_errors.mli @@ -14,14 +14,14 @@ open Type_errors (** {6 The type of errors raised by the pretyper } *) type unification_error = - | OccurCheck of existential_key * constr - | NotClean of existential * env * constr + | OccurCheck of existential_key * EConstr.constr + | NotClean of EConstr.existential * env * EConstr.constr | NotSameArgSize | NotSameHead | NoCanonicalStructure - | ConversionFailed of env * constr * constr + | ConversionFailed of env * EConstr.constr * EConstr.constr | MetaOccurInBody of existential_key - | InstanceNotSameType of existential_key * env * types * types + | InstanceNotSameType of existential_key * env * EConstr.types * EConstr.types | UnifUnivInconsistency of Univ.univ_inconsistency | CannotSolveConstraint of Evd.evar_constraint * unification_error | ProblemBeyondCapabilities diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml index b689bb7c7..fbba682fc 100644 --- a/pretyping/pretyping.ml +++ b/pretyping/pretyping.ml @@ -756,7 +756,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre refreshed right away. *) let sigma = !evdref in let c = mkApp (f,Array.map (whd_evar sigma) args) in - let c = evd_comb1 (Evarsolve.refresh_universes (Some true) env.ExtraEnv.env) evdref c in + let c = evd_comb1 (Evarsolve.refresh_universes (Some true) env.ExtraEnv.env) evdref (EConstr.of_constr c) in let t = Retyping.get_type_of env.ExtraEnv.env !evdref (EConstr.of_constr c) in make_judge c (* use this for keeping evars: resj.uj_val *) t else resj @@ -820,7 +820,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre in let t = evd_comb1 (Evarsolve.refresh_universes ~onlyalg:true ~status:Evd.univ_flexible (Some false) env.ExtraEnv.env) - evdref j.uj_type in + evdref (EConstr.of_constr j.uj_type) in (* The name specified by ltac is used also to create bindings. So the substitution must also be applied on variables before they are looked up in the rel context. *) @@ -1003,7 +1003,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre let tj = pretype_type empty_valcon env evdref lvar t in let tval = evd_comb1 (Evarsolve.refresh_universes ~onlyalg:true ~status:Evd.univ_flexible (Some false) env.ExtraEnv.env) - evdref tj.utj_val in + evdref (EConstr.of_constr tj.utj_val) in let tval = nf_evar !evdref tval in let cj, tval = match k with | VMcast -> @@ -1014,7 +1014,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre if b then (evdref := evd; cj, tval) else error_actual_type ~loc env.ExtraEnv.env !evdref cj tval - (ConversionFailed (env.ExtraEnv.env,cty,tval)) + (ConversionFailed (env.ExtraEnv.env,EConstr.of_constr cty,EConstr.of_constr tval)) else user_err ~loc (str "Cannot check cast with vm: " ++ str "unresolved arguments remain.") | NATIVEcast -> @@ -1025,7 +1025,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre if b then (evdref := evd; cj, tval) else error_actual_type ~loc env.ExtraEnv.env !evdref cj tval - (ConversionFailed (env.ExtraEnv.env,cty,tval)) + (ConversionFailed (env.ExtraEnv.env,EConstr.of_constr cty,EConstr.of_constr tval)) end | _ -> pretype (mk_tycon (EConstr.of_constr tval)) env evdref lvar c, tval diff --git a/pretyping/typing.ml b/pretyping/typing.ml index db31541cd..1dcb5f945 100644 --- a/pretyping/typing.ml +++ b/pretyping/typing.ml @@ -309,7 +309,7 @@ let type_of ?(refresh=false) env evd c = let j = execute env evdref c in (* side-effect on evdref *) if refresh then - Evarsolve.refresh_universes ~onlyalg:true (Some false) env !evdref j.uj_type + Evarsolve.refresh_universes ~onlyalg:true (Some false) env !evdref (EConstr.of_constr j.uj_type) else !evdref, j.uj_type let e_type_of ?(refresh=false) env evdref c = @@ -317,7 +317,7 @@ let e_type_of ?(refresh=false) env evdref c = let j = execute env evdref c in (* side-effect on evdref *) if refresh then - let evd, c = Evarsolve.refresh_universes ~onlyalg:true (Some false) env !evdref j.uj_type in + let evd, c = Evarsolve.refresh_universes ~onlyalg:true (Some false) env !evdref (EConstr.of_constr j.uj_type) in let () = evdref := evd in c else j.uj_type diff --git a/pretyping/unification.ml b/pretyping/unification.ml index f282ec4f1..b8c9a93db 100644 --- a/pretyping/unification.ml +++ b/pretyping/unification.ml @@ -183,7 +183,7 @@ let solve_pattern_eqn_array (env,nb) f l c (sigma,metasubst,evarsubst) = extra assumptions added by unification to the context *) let env' = pop_rel_context nb env in let sigma,c = pose_all_metas_as_evars env' sigma c in - let c = solve_pattern_eqn env sigma l c in + let c = solve_pattern_eqn env sigma (List.map EConstr.of_constr l) (EConstr.of_constr c) in let pb = (Conv,TypeNotProcessed) in if noccur_between 1 nb c then sigma,(k,lift (-nb) c,pb)::metasubst,evarsubst @@ -191,7 +191,7 @@ let solve_pattern_eqn_array (env,nb) f l c (sigma,metasubst,evarsubst) = | Evar ev -> let env' = pop_rel_context nb env in let sigma,c = pose_all_metas_as_evars env' sigma c in - sigma,metasubst,(env,ev,solve_pattern_eqn env sigma l c)::evarsubst + sigma,metasubst,(env,ev,solve_pattern_eqn env sigma (List.map EConstr.of_constr l) (EConstr.of_constr c))::evarsubst | _ -> assert false let push d (env,n) = (push_rel_assum d env,n+1) @@ -841,7 +841,7 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb and unify_app_pattern dir curenvnb pb opt substn cM f1 l1 cN f2 l2 = let f, l, t = if dir then f1, l1, cN else f2, l2, cM in - match is_unification_pattern curenvnb sigma f (Array.to_list l) t with + match is_unification_pattern curenvnb sigma (EConstr.of_constr f) (Array.map_to_list EConstr.of_constr l) (EConstr.of_constr t) with | None -> (match kind_of_term t with | App (f',l') -> @@ -850,7 +850,7 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb | Proj _ -> unify_app curenvnb pb opt substn cM f1 l1 cN f2 l2 | _ -> unify_not_same_head curenvnb pb opt substn cM cN) | Some l -> - solve_pattern_eqn_array curenvnb f l t substn + solve_pattern_eqn_array curenvnb f (List.map EConstr.Unsafe.to_constr l) t substn and unify_app (curenv, nb as curenvnb) pb opt (sigma, metas, evars as substn) cM f1 l1 cN f2 l2 = try @@ -1246,7 +1246,7 @@ let w_coerce env evd mv c = w_coerce_to_type env evd c cty mvty let unify_to_type env sigma flags c status u = - let sigma, c = refresh_universes (Some false) env sigma c in + let sigma, c = refresh_universes (Some false) env sigma (EConstr.of_constr c) in let t = get_type_of env sigma (EConstr.of_constr (nf_meta sigma c)) in let t = nf_betaiota sigma (EConstr.of_constr (nf_meta sigma t)) in unify_0 env sigma CUMUL flags t u @@ -1271,10 +1271,13 @@ let order_metas metas = (* Solve an equation ?n[x1=u1..xn=un] = t where ?n is an evar *) +let to_conv_fun f = (); fun env sigma pb c1 c2 -> + f env sigma pb (EConstr.Unsafe.to_constr c1) (EConstr.Unsafe.to_constr c2) + let solve_simple_evar_eqn ts env evd ev rhs = - match solve_simple_eqn (Evarconv.evar_conv_x ts) env evd (None,ev,EConstr.of_constr rhs) with + match solve_simple_eqn (to_conv_fun (Evarconv.evar_conv_x ts)) env evd (None,ev,EConstr.of_constr rhs) with | UnifFailure (evd,reason) -> - error_cannot_unify env evd ~reason (mkEvar ev,rhs); + error_cannot_unify env evd ~reason (EConstr.Unsafe.to_constr (EConstr.mkEvar ev),rhs); | Success evd -> Evarconv.consider_remaining_unif_problems env evd @@ -1308,14 +1311,14 @@ let w_merge env with_types flags (evd,metas,evars) = else let evd' = let evd', rhs'' = pose_all_metas_as_evars curenv evd rhs' in - try solve_simple_evar_eqn flags.modulo_delta_types curenv evd' ev rhs'' + try solve_simple_evar_eqn flags.modulo_delta_types curenv evd' (fst ev, Array.map EConstr.of_constr (snd ev)) rhs'' with Retyping.RetypeError _ -> error_cannot_unify curenv evd' (mkEvar ev,rhs'') in w_merge_rec evd' metas evars' eqns | _ -> let evd', rhs'' = pose_all_metas_as_evars curenv evd rhs' in let evd' = - try solve_simple_evar_eqn flags.modulo_delta_types curenv evd' ev rhs'' + try solve_simple_evar_eqn flags.modulo_delta_types curenv evd' (fst ev, Array.map EConstr.of_constr (snd ev)) rhs'' with Retyping.RetypeError _ -> error_cannot_unify curenv evd' (mkEvar ev, rhs'') in w_merge_rec evd' metas evars' eqns diff --git a/tactics/tactics.ml b/tactics/tactics.ml index 22d01e401..1c10cdfea 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -788,7 +788,7 @@ let check_types env sigma mayneedglobalcheck deep newc origc = if deep then begin let t2 = Retyping.get_type_of env sigma (EConstr.of_constr origc) in let sigma, t2 = Evarsolve.refresh_universes - ~onlyalg:true (Some false) env sigma t2 in + ~onlyalg:true (Some false) env sigma (EConstr.of_constr t2) in let sigma, b = infer_conv ~pb:Reduction.CUMUL env sigma t1 t2 in if not b then if @@ -2604,7 +2604,7 @@ let letin_tac_gen with_eq (id,depdecls,lastlhyp,ccl,c) ty = let Sigma (t, sigma, p) = match ty with | Some t -> Sigma.here t sigma | None -> - let t = typ_of env sigma (EConstr.of_constr c) in + let t = EConstr.of_constr (typ_of env sigma (EConstr.of_constr c)) in let sigma, c = Evarsolve.refresh_universes ~onlyalg:true (Some false) env (Sigma.to_evar_map sigma) t in Sigma.Unsafe.of_pair (c, sigma) in @@ -3621,7 +3621,7 @@ let abstract_args gl generalize_vars dep id defined f args = RelDecl.get_name decl, RelDecl.get_type decl, c in let argty = Tacmach.pf_unsafe_type_of gl arg in - let sigma', ty = Evarsolve.refresh_universes (Some true) env !sigma ty in + let sigma', ty = Evarsolve.refresh_universes (Some true) env !sigma (EConstr.of_constr ty) in let () = sigma := sigma' in let lenctx = List.length ctx in let liftargty = lift lenctx argty in diff --git a/toplevel/himsg.ml b/toplevel/himsg.ml index c09bf59ce..b480fcd86 100644 --- a/toplevel/himsg.ml +++ b/toplevel/himsg.ml @@ -60,13 +60,19 @@ let contract1_vect env a v = | _ -> assert false let rec contract3' env a b c = function - | OccurCheck (evk,d) -> let x,d = contract4 env a b c d in x,OccurCheck(evk,d) + | OccurCheck (evk,d) -> + let d = EConstr.Unsafe.to_constr d in + let x,d = contract4 env a b c d in x,OccurCheck(evk, EConstr.of_constr d) | NotClean ((evk,args),env',d) -> + let d = EConstr.Unsafe.to_constr d in + let args = Array.map EConstr.Unsafe.to_constr args in let env',d,args = contract1_vect env' d args in - contract3 env a b c,NotClean((evk,args),env',d) + contract3 env a b c,NotClean((evk,Array.map EConstr.of_constr args),env',EConstr.of_constr d) | ConversionFailed (env',t1,t2) -> + let t1 = EConstr.Unsafe.to_constr t1 in + let t2 = EConstr.Unsafe.to_constr t2 in let (env',t1,t2) = contract2 env' t1 t2 in - contract3 env a b c, ConversionFailed (env',t1,t2) + contract3 env a b c, ConversionFailed (env',EConstr.of_constr t1,EConstr.of_constr t2) | NotSameArgSize | NotSameHead | NoCanonicalStructure | MetaOccurInBody _ | InstanceNotSameType _ | ProblemBeyondCapabilities | UnifUnivInconsistency _ as x -> contract3 env a b c, x @@ -277,13 +283,13 @@ let explain_unification_error env sigma p1 p2 = function | Some e -> let rec aux p1 p2 = function | OccurCheck (evk,rhs) -> - let rhs = Evarutil.nf_evar sigma rhs in + let rhs = EConstr.to_constr sigma rhs in [str "cannot define " ++ quote (pr_existential_key sigma evk) ++ strbrk " with term " ++ pr_lconstr_env env sigma rhs ++ strbrk " that would depend on itself"] | NotClean ((evk,args),env,c) -> - let c = Evarutil.nf_evar sigma c in - let args = Array.map (Evarutil.nf_evar sigma) args in + let c = EConstr.to_constr sigma c in + let args = Array.map (EConstr.to_constr sigma) args in [str "cannot instantiate " ++ quote (pr_existential_key sigma evk) ++ strbrk " because " ++ pr_lconstr_env env sigma c ++ strbrk " is not in its scope" ++ @@ -293,6 +299,8 @@ let explain_unification_error env sigma p1 p2 = function | NotSameArgSize | NotSameHead | NoCanonicalStructure -> (* Error speaks from itself *) [] | ConversionFailed (env,t1,t2) -> + let t1 = EConstr.to_constr sigma t1 in + let t2 = EConstr.to_constr sigma t2 in if Term.eq_constr t1 p1 && Term.eq_constr t2 p2 then [] else let env = make_all_name_different env in let t1 = Evarutil.nf_evar sigma t1 in @@ -306,6 +314,8 @@ let explain_unification_error env sigma p1 p2 = function strbrk " refers to a metavariable - please report your example" ++ strbrk "at " ++ str Coq_config.wwwbugtracker ++ str "."] | InstanceNotSameType (evk,env,t,u) -> + let t = EConstr.to_constr sigma t in + let u = EConstr.to_constr sigma u in let t, u = pr_explicit env sigma t u in [str "unable to find a well-typed instantiation for " ++ quote (pr_existential_key sigma evk) ++ -- cgit v1.2.3 From b113f9e1ca88514cd9d94dfe90669a27689b7434 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sun, 6 Nov 2016 03:11:44 +0100 Subject: Recordops API using EConstr. --- pretyping/recordops.ml | 6 +++--- pretyping/recordops.mli | 2 +- 2 files changed, 4 insertions(+), 4 deletions(-) diff --git a/pretyping/recordops.ml b/pretyping/recordops.ml index 062e4a068..f09f3221d 100644 --- a/pretyping/recordops.ml +++ b/pretyping/recordops.ml @@ -171,7 +171,7 @@ let keep_true_projections projs kinds = let filter (p, (_, b)) = if b then Some p else None in List.map_filter filter (List.combine projs kinds) -let cs_pattern_of_constr sigma t = +let cs_pattern_of_constr t = match kind_of_term t with App (f,vargs) -> begin @@ -179,7 +179,7 @@ let cs_pattern_of_constr sigma t = with e when CErrors.noncritical e -> raise Not_found end | Rel n -> Default_cs, Some n, [] - | Prod (_,a,b) when EConstr.Vars.noccurn sigma 1 (EConstr.of_constr b) -> Prod_cs, None, [a; Termops.pop (EConstr.of_constr b)] + | Prod (_,a,b) when Vars.noccurn 1 b -> Prod_cs, None, [a; Vars.lift (-1) b] | Sort s -> Sort_cs (family_of_sort s), None, [] | _ -> begin @@ -217,7 +217,7 @@ let compute_canonical_projections warn (con,ind) = | Some proji_sp -> begin try - let patt, n , args = cs_pattern_of_constr Evd.empty t (** FIXME *) in + let patt, n , args = cs_pattern_of_constr t in ((ConstRef proji_sp, patt, t, n, args) :: l) with Not_found -> let con_pp = Nametab.pr_global_env Id.Set.empty (ConstRef con) diff --git a/pretyping/recordops.mli b/pretyping/recordops.mli index 405963a9c..7c0d0ec6d 100644 --- a/pretyping/recordops.mli +++ b/pretyping/recordops.mli @@ -65,7 +65,7 @@ type obj_typ = { o_TCOMPS : constr list } (** ordered *) (** Return the form of the component of a canonical structure *) -val cs_pattern_of_constr : Evd.evar_map -> constr -> cs_pattern * int option * constr list +val cs_pattern_of_constr : constr -> cs_pattern * int option * constr list val pr_cs_pattern : cs_pattern -> Pp.std_ppcmds -- cgit v1.2.3 From b365304d32db443194b7eaadda63c784814f53f1 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sun, 6 Nov 2016 03:23:13 +0100 Subject: Evarconv API using EConstr. --- engine/evarutil.ml | 37 +-- engine/evarutil.mli | 6 +- engine/termops.ml | 11 + engine/termops.mli | 1 + engine/universes.ml | 9 - engine/universes.mli | 2 - ltac/rewrite.ml | 2 +- plugins/funind/functional_principles_proofs.ml | 2 +- plugins/ssrmatching/ssrmatching.ml4 | 4 +- pretyping/cases.ml | 13 +- pretyping/coercion.ml | 12 +- pretyping/evarconv.ml | 342 +++++++++++++------------ pretyping/evarconv.mli | 9 +- pretyping/evardefine.ml | 2 +- pretyping/evardefine.mli | 2 +- pretyping/pretyping.ml | 6 +- pretyping/typeclasses.ml | 2 +- pretyping/typing.ml | 16 +- pretyping/unification.ml | 23 +- proofs/evar_refiner.ml | 2 +- proofs/pfedit.ml | 2 +- stm/stm.ml | 7 +- tactics/class_tactics.ml | 4 +- tactics/equality.ml | 6 +- tactics/tactics.ml | 4 +- toplevel/vernacentries.ml | 4 +- 26 files changed, 255 insertions(+), 275 deletions(-) diff --git a/engine/evarutil.ml b/engine/evarutil.ml index 05f98a41f..62627a416 100644 --- a/engine/evarutil.ml +++ b/engine/evarutil.ml @@ -148,21 +148,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 @@ -171,10 +161,10 @@ let is_ground_term evd t = let is_ground_env evd env = let is_ground_rel_decl = function - | RelDecl.LocalDef (_,b,_) -> is_ground_term evd b + | RelDecl.LocalDef (_,b,_) -> is_ground_term evd (EConstr.of_constr b) | _ -> true in let is_ground_named_decl = function - | NamedDecl.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) @@ -208,27 +198,18 @@ let head_evar = 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) diff --git a/engine/evarutil.mli b/engine/evarutil.mli index a2e2a07dd..d6e2d4534 100644 --- a/engine/evarutil.mli +++ b/engine/evarutil.mli @@ -91,12 +91,12 @@ exception NoHeadEvar val head_evar : constr -> existential_key (** may raise NoHeadEvar *) (* Expand head evar if any *) -val whd_head_evar : evar_map -> constr -> constr +val whd_head_evar : evar_map -> EConstr.constr -> EConstr.constr (* An over-approximation of [has_undefined (nf_evars evd c)] *) -val has_undefined_evars : evar_map -> constr -> bool +val has_undefined_evars : evar_map -> EConstr.constr -> bool -val is_ground_term : evar_map -> constr -> bool +val is_ground_term : evar_map -> EConstr.constr -> bool val is_ground_env : evar_map -> env -> bool (** [gather_dependent_evars evm seeds] classifies the evars in [evm] diff --git a/engine/termops.ml b/engine/termops.ml index 356312e2f..26b22be4e 100644 --- a/engine/termops.ml +++ b/engine/termops.ml @@ -1068,6 +1068,17 @@ let dependency_closure env sigma sign hyps = sign in List.rev lh +let global_app_of_constr sigma c = + let open Univ in + let open Globnames in + match EConstr.kind sigma c with + | Const (c, u) -> (ConstRef c, u), None + | Ind (i, u) -> (IndRef i, u), None + | Construct (c, u) -> (ConstructRef c, u), None + | Var id -> (VarRef id, Instance.empty), None + | Proj (p, c) -> (ConstRef (Projection.constant p), Instance.empty), Some c + | _ -> raise Not_found + (* Combinators on judgments *) let on_judgment f j = { uj_val = f j.uj_val; uj_type = f j.uj_type } diff --git a/engine/termops.mli b/engine/termops.mli index b536b0fb8..24c2c82f2 100644 --- a/engine/termops.mli +++ b/engine/termops.mli @@ -254,6 +254,7 @@ val clear_named_body : Id.t -> env -> env val global_vars : env -> Evd.evar_map -> EConstr.t -> Id.t list val global_vars_set_of_decl : env -> Evd.evar_map -> Context.Named.Declaration.t -> Id.Set.t +val global_app_of_constr : Evd.evar_map -> EConstr.constr -> Globnames.global_reference puniverses * EConstr.constr option (** Gives an ordered list of hypotheses, closed by dependencies, containing a given set *) diff --git a/engine/universes.ml b/engine/universes.ml index 6720fcef8..51a113219 100644 --- a/engine/universes.ml +++ b/engine/universes.ml @@ -446,15 +446,6 @@ let global_of_constr c = | Var id -> VarRef id, Instance.empty | _ -> raise Not_found -let global_app_of_constr c = - match kind_of_term 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, Instance.empty), None - | Proj (p, c) -> (ConstRef (Projection.constant p), Instance.empty), Some c - | _ -> raise Not_found - open Declarations let type_of_reference env r = diff --git a/engine/universes.mli b/engine/universes.mli index 725c21d29..c3e2055f3 100644 --- a/engine/universes.mli +++ b/engine/universes.mli @@ -130,8 +130,6 @@ val fresh_universe_context_set_instance : universe_context_set -> (** Raises [Not_found] if not a global reference. *) val global_of_constr : constr -> Globnames.global_reference puniverses -val global_app_of_constr : constr -> Globnames.global_reference puniverses * constr option - val constr_of_global_univ : Globnames.global_reference puniverses -> constr val extend_context : 'a in_universe_context_set -> universe_context_set -> diff --git a/ltac/rewrite.ml b/ltac/rewrite.ml index 80307ce8e..076e4c05e 100644 --- a/ltac/rewrite.ml +++ b/ltac/rewrite.ml @@ -454,7 +454,7 @@ let evd_convertible env evd x y = unsolvable constraints remain, so we check that this unification does not introduce any new problem. *) let _, pbs = Evd.extract_all_conv_pbs evd in - let evd' = Evarconv.the_conv_x env x y evd in + let evd' = Evarconv.the_conv_x env (EConstr.of_constr x) (EConstr.of_constr y) evd in let _, pbs' = Evd.extract_all_conv_pbs evd' in if evd' == evd || problem_inclusion pbs' pbs then Some evd' else None diff --git a/plugins/funind/functional_principles_proofs.ml b/plugins/funind/functional_principles_proofs.ml index 258ee5ad6..340dd2c28 100644 --- a/plugins/funind/functional_principles_proofs.ml +++ b/plugins/funind/functional_principles_proofs.ml @@ -237,7 +237,7 @@ let change_eq env sigma hyp_id (context:Context.Rel.t) x t end_of_type = failwith "NoChange"; end in - let eq_constr = Evarconv.e_conv env (ref sigma) in + let eq_constr c1 c2 = Evarconv.e_conv env (ref sigma) (EConstr.of_constr c1) (EConstr.of_constr c2) in if not (noccurn 1 end_of_type) then nochange "dependent"; (* if end_of_type depends on this term we don't touch it *) if not (isApp t) then nochange "not an equality"; diff --git a/plugins/ssrmatching/ssrmatching.ml4 b/plugins/ssrmatching/ssrmatching.ml4 index cc39b7260..e0d99d453 100644 --- a/plugins/ssrmatching/ssrmatching.ml4 +++ b/plugins/ssrmatching/ssrmatching.ml4 @@ -298,9 +298,9 @@ let unif_EQ_args env sigma pa a = prof_unif_eq_args.profile (unif_EQ_args env sigma pa) a ;; -let unif_HO env ise p c = Evarconv.the_conv_x env p c ise +let unif_HO env ise p c = Evarconv.the_conv_x env (EConstr.of_constr p) (EConstr.of_constr c) ise -let unif_HOtype env ise p c = Evarconv.the_conv_x_leq env p c ise +let unif_HOtype env ise p c = Evarconv.the_conv_x_leq env (EConstr.of_constr p) (EConstr.of_constr c) ise let unif_HO_args env ise0 pa i ca = let n = Array.length pa in diff --git a/pretyping/cases.ml b/pretyping/cases.ml index a68daf4e5..04f50d50e 100644 --- a/pretyping/cases.ml +++ b/pretyping/cases.ml @@ -78,9 +78,6 @@ let list_try_compile f l = let force_name = let nx = Name default_dependent_ident in function Anonymous -> nx | na -> na -let to_conv_fun f = (); fun env sigma pb c1 c2 -> - f env sigma pb (EConstr.Unsafe.to_constr c1) (EConstr.Unsafe.to_constr c2) - (************************************************************************) (* Pattern-matching compilation (Cases) *) (************************************************************************) @@ -311,7 +308,7 @@ let inh_coerce_to_ind evdref env loc ty tyi = constructor and renounce if not able to give more information *) (* devrait être indifférent d'exiger leq ou pas puisque pour un inductif cela doit être égal *) - if not (e_cumul env evdref expected_typ ty) then evdref := sigma + if not (e_cumul env evdref (EConstr.of_constr expected_typ) (EConstr.of_constr ty)) then evdref := sigma let binding_vars_of_inductive = function | NotInd _ -> [] @@ -395,7 +392,7 @@ let adjust_tomatch_to_pattern pb ((current,typ),deps,dep) = let current = if List.is_empty deps && isEvar typ then (* Don't insert coercions if dependent; only solve evars *) - let _ = e_cumul pb.env pb.evdref indt typ in + let _ = e_cumul pb.env pb.evdref (EConstr.of_constr indt) (EConstr.of_constr typ) in current else (evd_comb2 (Coercion.inh_conv_coerce_to true Loc.ghost pb.env) @@ -1639,7 +1636,7 @@ let abstract_tycon loc env evdref subst tycon extenv t = 1 (rel_context env) in let ev' = e_new_evar env evdref ~src ty in let ev = (fst ev, Array.map EConstr.of_constr (snd ev)) in - begin match solve_simple_eqn (to_conv_fun (evar_conv_x full_transparent_state)) env !evdref (None,ev,EConstr.of_constr (substl inst ev')) with + begin match solve_simple_eqn (evar_conv_x full_transparent_state) env !evdref (None,ev,EConstr.of_constr (substl inst ev')) with | Success evd -> evdref := evd | UnifFailure _ -> assert false end; @@ -1690,7 +1687,7 @@ let build_tycon loc env tycon_env s subst tycon extenv evdref t = let evd,tt = Typing.type_of extenv !evdref t in evdref := evd; (t,tt) in - let b = e_cumul env evdref tt (mkSort s) (* side effect *) in + let b = e_cumul env evdref (EConstr.of_constr tt) (EConstr.mkSort s) (* side effect *) in if not b then anomaly (Pp.str "Build_tycon: should be a type"); { uj_val = t; uj_type = tt } @@ -2083,7 +2080,7 @@ let constr_of_pat env evdref arsign pat avoid = try let env = push_rel_context sign env in evdref := the_conv_x_leq (push_rel_context sign env) - (lift (succ m) ty) (lift 1 apptype) !evdref; + (EConstr.of_constr (lift (succ m) ty)) (EConstr.of_constr (lift 1 apptype)) !evdref; let eq_t = mk_eq evdref (lift (succ m) ty) (mkRel 1) (* alias *) (lift 1 app) (* aliased term *) diff --git a/pretyping/coercion.ml b/pretyping/coercion.ml index c5418d22e..0ea6758a7 100644 --- a/pretyping/coercion.ml +++ b/pretyping/coercion.ml @@ -65,7 +65,7 @@ let apply_coercion_args env evd check isproj argl funj = | h::restl -> (* On devrait pouvoir s'arranger pour qu'on n'ait pas a faire hnf_constr *) match kind_of_term (whd_all env evd (EConstr.of_constr typ)) with | Prod (_,c1,c2) -> - if check && not (e_cumul env evdref (Retyping.get_type_of env evd (EConstr.of_constr h)) c1) then + if check && not (e_cumul env evdref (EConstr.of_constr (Retyping.get_type_of env evd (EConstr.of_constr h))) (EConstr.of_constr c1)) then raise NoCoercion; apply_rec (h::acc) (subst1 h c2) restl | _ -> anomaly (Pp.str "apply_coercion_args") @@ -147,7 +147,7 @@ and coerce loc env evdref (x : Term.constr) (y : Term.constr) let rec coerce_unify env x y = let x = hnf env !evdref (EConstr.of_constr x) and y = hnf env !evdref (EConstr.of_constr y) in try - evdref := the_conv_x_leq env x y !evdref; + evdref := the_conv_x_leq env (EConstr.of_constr x) (EConstr.of_constr y) !evdref; None with UnableToUnify _ -> coerce' env x y and coerce' env x y : (Term.constr -> Term.constr) option = @@ -162,7 +162,7 @@ and coerce loc env evdref (x : Term.constr) (y : Term.constr) let rec aux tele typ typ' i co = if i < len then let hdx = l.(i) and hdy = l'.(i) in - try evdref := the_conv_x_leq env hdx hdy !evdref; + try evdref := the_conv_x_leq env (EConstr.of_constr hdx) (EConstr.of_constr hdy) !evdref; let (n, eqT), restT = dest_prod typ in let (n', eqT'), restT' = dest_prod typ' in aux (hdx :: tele) (subst1 hdx restT) (subst1 hdy restT') (succ i) co @@ -170,7 +170,7 @@ and coerce loc env evdref (x : Term.constr) (y : Term.constr) let (n, eqT), restT = dest_prod typ in let (n', eqT'), restT' = dest_prod typ' in let _ = - try evdref := the_conv_x_leq env eqT eqT' !evdref + try evdref := the_conv_x_leq env (EConstr.of_constr eqT) (EConstr.of_constr eqT') !evdref with UnableToUnify _ -> raise NoSubtacCoercion in (* Disallow equalities on arities *) @@ -458,11 +458,11 @@ let inh_coerce_to_fail env evd rigidonly v t c1 = | None -> evd, None, t with Not_found -> raise NoCoercion in - try (the_conv_x_leq env t' c1 evd, v') + try (the_conv_x_leq env (EConstr.of_constr t') (EConstr.of_constr c1) evd, v') with UnableToUnify _ -> raise NoCoercion let rec inh_conv_coerce_to_fail loc env evd rigidonly v t c1 = - try (the_conv_x_leq env t c1 evd, v) + try (the_conv_x_leq env (EConstr.of_constr t) (EConstr.of_constr c1) evd, v) with UnableToUnify (best_failed_evd,e) -> try inh_coerce_to_fail env evd rigidonly v t c1 with NoCoercion -> diff --git a/pretyping/evarconv.ml b/pretyping/evarconv.ml index 8f3f671ab..c8dcb19b4 100644 --- a/pretyping/evarconv.ml +++ b/pretyping/evarconv.ml @@ -30,7 +30,7 @@ module RelDecl = Context.Rel.Declaration module NamedDecl = Context.Named.Declaration type unify_fun = transparent_state -> - env -> evar_map -> conv_pb -> constr -> constr -> Evarsolve.unification_result + env -> evar_map -> conv_pb -> EConstr.constr -> EConstr.constr -> Evarsolve.unification_result let debug_unification = ref (false) let _ = Goptions.declare_bool_option { @@ -42,33 +42,32 @@ let _ = Goptions.declare_bool_option { Goptions.optwrite = (fun a -> debug_unification:=a); } -let to_conv_fun f = (); fun env sigma pb c1 c2 -> - f env sigma pb (EConstr.Unsafe.to_constr c1) (EConstr.Unsafe.to_constr c2) - let unfold_projection env evd ts p c = + let open EConstr in let cst = Projection.constant p in if is_transparent_constant ts cst then Some (mkProj (Projection.make cst true, c)) else None let eval_flexible_term ts env evd c = - match kind_of_term c with + let open EConstr in + match EConstr.kind evd c with | Const (c,u as cu) -> if is_transparent_constant ts c - then constant_opt_value_in env cu + then Option.map EConstr.of_constr (constant_opt_value_in env cu) else None | Rel n -> (try match lookup_rel n env with | RelDecl.LocalAssum _ -> None - | RelDecl.LocalDef (_,v,_) -> Some (lift n v) + | RelDecl.LocalDef (_,v,_) -> Some (Vars.lift n (EConstr.of_constr v)) with Not_found -> None) | Var id -> (try if is_transparent_variable ts id then - env |> lookup_named id |> NamedDecl.get_value + Option.map EConstr.of_constr (env |> lookup_named id |> NamedDecl.get_value) else None with Not_found -> None) - | LetIn (_,b,_,c) -> Some (subst1 b c) + | LetIn (_,b,_,c) -> Some (Vars.subst1 b c) | Lambda _ -> Some c | Proj (p, c) -> if Projection.unfolded p then assert false @@ -77,12 +76,11 @@ let eval_flexible_term ts env evd c = type flex_kind_of_term = | Rigid - | MaybeFlexible of Constr.t (* reducible but not necessarily reduced *) - | Flexible of existential + | MaybeFlexible of EConstr.t (* reducible but not necessarily reduced *) + | Flexible of EConstr.existential let flex_kind_of_term ts env evd c sk = - let c = EConstr.Unsafe.to_constr c in - match kind_of_term c with + match EConstr.kind evd c with | LetIn _ | Rel _ | Const _ | Var _ | Proj _ -> Option.cata (fun x -> MaybeFlexible x) Rigid (eval_flexible_term ts env evd c) | Lambda _ when not (Option.is_empty (Stack.decomp sk)) -> MaybeFlexible c @@ -92,12 +90,13 @@ let flex_kind_of_term ts env evd c sk = | Fix _ -> Rigid (* happens when the fixpoint is partially applied *) | Cast _ | App _ | Case _ -> assert false -let zip evd (c, stk) = EConstr.Unsafe.to_constr (Stack.zip evd (c, stk)) +let add_conv_pb (pb, env, x, y) sigma = + Evd.add_conv_pb (pb, env, EConstr.Unsafe.to_constr x, EConstr.Unsafe.to_constr y) sigma let apprec_nohdbeta ts env evd c = - let (t,sk as appr) = Reductionops.whd_nored_state evd (EConstr.of_constr c, []) in + let (t,sk as appr) = Reductionops.whd_nored_state evd (c, []) in if Stack.not_purely_applicative sk - then zip evd (fst (whd_betaiota_deltazeta_for_iota_state + then Stack.zip evd (fst (whd_betaiota_deltazeta_for_iota_state ts env evd Cst_stack.empty appr)) else c @@ -106,8 +105,9 @@ let position_problem l2r = function | CUMUL -> Some l2r let occur_rigidly (evk,_ as ev) evd t = + let open EConstr in let rec aux t = - match kind_of_term (whd_evar evd t) with + match EConstr.kind evd t with | App (f, c) -> if aux f then Array.exists aux c else false | Construct _ | Ind _ | Sort _ | Meta _ | Fix _ | CoFix _ -> true | Proj (p, c) -> not (aux c) @@ -117,7 +117,7 @@ let occur_rigidly (evk,_ as ev) evd t = | Const _ -> false | Prod (_, b, t) -> ignore(aux b || aux t); true | Rel _ | Var _ -> false - | Case (_,_,c,_) -> if eq_constr (mkEvar ev) c then raise Occur else false + | Case (_,_,c,_) -> if eq_constr evd (mkEvar ev) c then raise Occur else false in try ignore(aux t); false with Occur -> true (* [check_conv_record env sigma (t1,stack1) (t2,stack2)] tries to decompose @@ -141,23 +141,22 @@ let occur_rigidly (evk,_ as ev) evd t = projection would have been reduced) *) let check_conv_record env sigma (t1,sk1) (t2,sk2) = - let t1 = EConstr.Unsafe.to_constr t1 in - let t2 = EConstr.Unsafe.to_constr t2 in - let (proji, u), arg = Universes.global_app_of_constr t1 in + let open EConstr in + let (proji, u), arg = Termops.global_app_of_constr sigma t1 in let canon_s,sk2_effective = try - match kind_of_term t2 with + match EConstr.kind sigma t2 with Prod (_,a,b) -> (* assert (l2=[]); *) - let _, a, b = destProd (Evarutil.nf_evar sigma t2) in - if EConstr.Vars.noccurn sigma 1 (EConstr.of_constr b) then + let _, a, b = destProd sigma t2 in + if Vars.noccurn sigma 1 b then lookup_canonical_conversion (proji, Prod_cs), - (Stack.append_app [|EConstr.of_constr a;EConstr.of_constr (pop (EConstr.of_constr b))|] Stack.empty) + (Stack.append_app [|a;EConstr.of_constr (pop b)|] Stack.empty) else raise Not_found | Sort s -> lookup_canonical_conversion (proji, Sort_cs (family_of_sort s)),[] | _ -> - let c2 = global_of_constr t2 in + let c2 = global_of_constr (EConstr.to_constr sigma t2) in lookup_canonical_conversion (proji, Const_cs c2),sk2 with Not_found -> let (c, cs) = lookup_canonical_conversion (proji,Default_cs) in @@ -165,17 +164,19 @@ let check_conv_record env sigma (t1,sk1) (t2,sk2) = in let t', { o_DEF = c; o_CTX = ctx; o_INJ=n; o_TABS = bs; o_TPARAMS = params; o_NPARAMS = nparams; o_TCOMPS = us } = canon_s in + let us = List.map EConstr.of_constr us in + let params = List.map EConstr.of_constr params in let params1, c1, extra_args1 = match arg with | Some c -> (* A primitive projection applied to c *) - let ty = Retyping.get_type_of ~lax:true env sigma (EConstr.of_constr c) in + let ty = Retyping.get_type_of ~lax:true env sigma c in let (i,u), ind_args = try Inductiveops.find_mrectype env sigma (EConstr.of_constr ty) with _ -> raise Not_found in Stack.append_app_list (List.map EConstr.of_constr ind_args) Stack.empty, c, sk1 | None -> match Stack.strip_n_app nparams sk1 with - | Some (params1, c1, extra_args1) -> params1, EConstr.Unsafe.to_constr c1, extra_args1 + | Some (params1, c1, extra_args1) -> params1, c1, extra_args1 | _ -> raise Not_found in let us2,extra_args2 = let l_us = List.length us in @@ -184,13 +185,13 @@ let check_conv_record env sigma (t1,sk1) (t2,sk2) = | None -> raise Not_found | Some (l',el,s') -> (l'@Stack.append_app [|el|] Stack.empty,s') in let subst, ctx' = Universes.fresh_universe_context_set_instance ctx in - let c' = subst_univs_level_constr subst c in + let c' = EConstr.of_constr (subst_univs_level_constr subst c) in let t' = subst_univs_level_constr subst t' in - let bs' = List.map (subst_univs_level_constr subst) bs in + let bs' = List.map (subst_univs_level_constr subst %> EConstr.of_constr) bs in let h, _ = decompose_app_vect sigma (EConstr.of_constr t') in - ctx',(h, t2),c',bs',(Stack.append_app_list (List.map EConstr.of_constr params) Stack.empty,params1), - (Stack.append_app_list (List.map EConstr.of_constr us) Stack.empty,us2),(extra_args1,extra_args2),c1, - (n, zip sigma (EConstr.of_constr t2,sk2)) + ctx',(EConstr.of_constr h, t2),c',bs',(Stack.append_app_list params Stack.empty,params1), + (Stack.append_app_list us Stack.empty,us2),(extra_args1,extra_args2),c1, + (n, Stack.zip sigma (t2,sk2)) (* Precondition: one of the terms of the pb is an uninstantiated evar, * possibly applied to arguments. *) @@ -220,11 +221,10 @@ let ise_exact ise x1 x2 = | Some _, Success i -> UnifFailure (i,NotSameArgSize) let ise_array2 evd f v1 v2 = - let inj c = EConstr.Unsafe.to_constr c in let rec allrec i = function | -1 -> Success i | n -> - match f i (inj v1.(n)) (inj v2.(n)) with + match f i v1.(n) v2.(n) with | Success i' -> allrec i' (n-1) | UnifFailure _ as x -> x in let lv1 = Array.length v1 in @@ -234,14 +234,13 @@ let ise_array2 evd f v1 v2 = (* Applicative node of stack are read from the outermost to the innermost but are unified the other way. *) let rec ise_app_stack2 env f evd sk1 sk2 = - let inj = EConstr.Unsafe.to_constr in match sk1,sk2 with | Stack.App node1 :: q1, Stack.App node2 :: q2 -> let (t1,l1) = Stack.decomp_node_last node1 q1 in let (t2,l2) = Stack.decomp_node_last node2 q2 in begin match ise_app_stack2 env f evd l1 l2 with |(_,UnifFailure _) as x -> x - |x,Success i' -> x,f env i' CONV (inj t1) (inj t2) + |x,Success i' -> x,f env i' CONV t1 t2 end | _, _ -> (sk1,sk2), Success evd @@ -255,14 +254,13 @@ let push_rec_types pfix env = stacks but not the entire stacks, the first part of the answer is Some(the remaining prefixes to tackle)) *) let ise_stack2 no_app env evd f sk1 sk2 = - let inj = EConstr.Unsafe.to_constr in let rec ise_stack2 deep i sk1 sk2 = let fail x = if deep then Some (List.rev sk1, List.rev sk2), Success i else None, x in match sk1, sk2 with | [], [] -> None, Success i | Stack.Case (_,t1,c1,_)::q1, Stack.Case (_,t2,c2,_)::q2 -> - (match f env i CONV (inj t1) (inj t2) with + (match f env i CONV t1 t2 with | Success i' -> (match ise_array2 i' (fun ii -> f env ii CONV) c1 c2 with | Success i'' -> ise_stack2 true i'' q1 q2 @@ -295,7 +293,6 @@ let ise_stack2 no_app env evd f sk1 sk2 = (* Make sure that the matching suffix is the all stack *) let exact_ise_stack2 env evd f sk1 sk2 = - let inj = EConstr.Unsafe.to_constr in let rec ise_stack2 i sk1 sk2 = match sk1, sk2 with | [], [] -> Success i @@ -303,7 +300,7 @@ let exact_ise_stack2 env evd f sk1 sk2 = ise_and i [ (fun i -> ise_stack2 i q1 q2); (fun i -> ise_array2 i (fun ii -> f env ii CONV) c1 c2); - (fun i -> f env i CONV (inj t1) (inj t2))] + (fun i -> f env i CONV t1 t2)] | Stack.Fix (((li1, i1),(_,tys1,bds1 as recdef1)),a1,_)::q1, Stack.Fix (((li2, i2),(_,tys2,bds2)),a2,_)::q2 -> if Int.equal i1 i2 && Array.equal Int.equal li1 li2 then @@ -341,10 +338,10 @@ let rec evar_conv_x ts env evd pbty term1 term2 = let e = try let evd, b = infer_conv ~catch_incon:false ~pb:pbty ~ts:(fst ts) - env evd term1 term2 + env evd (EConstr.Unsafe.to_constr term1) (EConstr.Unsafe.to_constr term2) in if b then Success evd - else UnifFailure (evd, ConversionFailed (env,EConstr.of_constr term1,EConstr.of_constr term2)) + else UnifFailure (evd, ConversionFailed (env,term1,term2)) with Univ.UniverseInconsistency e -> UnifFailure (evd, UnifUnivInconsistency e) in match e with @@ -361,19 +358,19 @@ let rec evar_conv_x ts env evd pbty term1 term2 = let term2 = apprec_nohdbeta (fst ts) env evd term2 in let default () = evar_eqappr_x ts env evd pbty - (whd_nored_state evd (EConstr.of_constr term1,Stack.empty), Cst_stack.empty) - (whd_nored_state evd (EConstr.of_constr term2,Stack.empty), Cst_stack.empty) + (whd_nored_state evd (term1,Stack.empty), Cst_stack.empty) + (whd_nored_state evd (term2,Stack.empty), Cst_stack.empty) in - begin match EConstr.kind evd (EConstr.of_constr term1), EConstr.kind evd (EConstr.of_constr term2) with + begin match EConstr.kind evd term1, EConstr.kind evd term2 with | Evar ev, _ when Evd.is_undefined evd (fst ev) -> - (match solve_simple_eqn (to_conv_fun (evar_conv_x ts)) env evd - (position_problem true pbty,ev, EConstr.of_constr term2) with + (match solve_simple_eqn (evar_conv_x ts) env evd + (position_problem true pbty,ev, term2) with | UnifFailure (_,OccurCheck _) -> (* Eta-expansion might apply *) default () | x -> x) | _, Evar ev when Evd.is_undefined evd (fst ev) -> - (match solve_simple_eqn (to_conv_fun (evar_conv_x ts)) env evd - (position_problem false pbty,ev, EConstr.of_constr term1) with + (match solve_simple_eqn (evar_conv_x ts) env evd + (position_problem false pbty,ev, term1) with | UnifFailure (_, OccurCheck _) -> (* Eta-expansion might apply *) default () | x -> x) @@ -382,6 +379,7 @@ let rec evar_conv_x ts env evd pbty term1 term2 = and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty ((term1,sk1 as appr1),csts1) ((term2,sk2 as appr2),csts2) = + let open EConstr in let quick_fail i = (* not costly, loses info *) UnifFailure (i, NotSameHead) in @@ -391,28 +389,27 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty | Some l1' -> (* Miller-Pfenning's patterns unification *) let t2 = EConstr.of_constr (nf_evar evd (EConstr.Unsafe.to_constr tM)) (** FIXME *) in let t2 = solve_pattern_eqn env evd l1' t2 in - solve_simple_eqn (to_conv_fun (evar_conv_x ts)) env evd - (position_problem on_left pbty,ev, EConstr.of_constr t2) + solve_simple_eqn (evar_conv_x ts) env evd + (position_problem on_left pbty,ev,EConstr.of_constr t2) in let consume_stack on_left (termF,skF) (termO,skO) evd = - let inj = EConstr.Unsafe.to_constr in let switch f a b = if on_left then f a b else f b a in let not_only_app = Stack.not_purely_applicative skO in match switch (ise_stack2 not_only_app env evd (evar_conv_x ts)) skF skO with |Some (l,r), Success i' when on_left && (not_only_app || List.is_empty l) -> - switch (evar_conv_x ts env i' pbty) (zip evd (termF,l)) (zip evd (termO,r)) + switch (evar_conv_x ts env i' pbty) (Stack.zip evd (termF,l)) (Stack.zip evd (termO,r)) |Some (r,l), Success i' when not on_left && (not_only_app || List.is_empty l) -> - switch (evar_conv_x ts env i' pbty) (zip evd (termF,l)) (zip evd (termO,r)) - |None, Success i' -> switch (evar_conv_x ts env i' pbty) (inj termF) (inj termO) + switch (evar_conv_x ts env i' pbty) (Stack.zip evd (termF,l)) (Stack.zip evd (termO,r)) + |None, Success i' -> switch (evar_conv_x ts env i' pbty) termF termO |_, (UnifFailure _ as x) -> x |Some _, _ -> UnifFailure (evd,NotSameArgSize) in let eta env evd onleft sk term sk' term' = assert (match sk with [] -> true | _ -> false); - let (na,c1,c'1) = destLambda term in - let c = nf_evar evd c1 in + let (na,c1,c'1) = destLambda evd term in + let c = EConstr.to_constr evd c1 in let env' = push_rel (RelDecl.LocalAssum (na,c)) env in let out1 = whd_betaiota_deltazeta_for_iota_state - (fst ts) env' evd Cst_stack.empty (EConstr.of_constr c'1, Stack.empty) in + (fst ts) env' evd Cst_stack.empty (c'1, Stack.empty) in let out2 = whd_nored_state evd (Stack.zip evd (term', sk' @ [Stack.Shift 1]), Stack.append_app [|EConstr.mkRel 1|] Stack.empty), Cst_stack.empty in @@ -438,12 +435,12 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty match Stack.list_of_app_stack skF with | None -> quick_fail evd | Some lF -> - let tM = zip evd apprM in + let tM = Stack.zip evd apprM in miller_pfenning on_left (fun () -> if not_only_app then (* Postpone the use of an heuristic *) - switch (fun x y -> Success (add_conv_pb (pbty,env,x,y) i)) (zip evd apprF) tM + switch (fun x y -> Success (add_conv_pb (pbty,env,x,y) i)) (Stack.zip evd apprF) tM else quick_fail i) - ev lF (EConstr.of_constr tM) i + ev lF tM i and consume (termF,skF as apprF) (termM,skM as apprM) i = if not (Stack.is_empty skF && Stack.is_empty skM) then consume_stack on_left apprF apprM i @@ -487,7 +484,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty let eta evd = match EConstr.kind evd termR with | Lambda _ when (* if ever problem is ill-typed: *) List.is_empty skR -> - eta env evd false skR (EConstr.Unsafe.to_constr termR) skF termF + eta env evd false skR termR skF termF | Construct u -> eta_constructor ts env evd skR u skF termF | _ -> UnifFailure (evd,NotSameHead) in @@ -495,7 +492,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty | None -> ise_try evd [consume_stack on_left apprF apprR; eta] | Some lF -> - let tR = zip evd apprR in + let tR = Stack.zip evd apprR in miller_pfenning on_left (fun () -> ise_try evd @@ -503,17 +500,17 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty (fun i -> if not (occur_rigidly ev i tR) then let i,tF = - if isRel tR || isVar tR then + if isRel i tR || isVar i tR then (* Optimization so as to generate candidates *) - let i,ev = evar_absorb_arguments env i (fst ev, Array.map EConstr.of_constr (snd ev)) lF in + let i,ev = evar_absorb_arguments env i ev lF in i,mkEvar ev else - i,zip evd apprF in + i,Stack.zip evd apprF in switch (fun x y -> Success (add_conv_pb (pbty,env,x,y) i)) tF tR else - UnifFailure (evd,OccurCheck (fst ev,EConstr.of_constr tR)))]) - (fst ev, Array.map EConstr.of_constr (snd ev)) lF (EConstr.of_constr tR) evd + UnifFailure (evd,OccurCheck (fst ev,tR)))]) + ev lF tR evd in let app_empty = match sk1, sk2 with [], [] -> true | _ -> false in (* Evar must be undefined since we have flushed evars *) @@ -531,20 +528,20 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty | None, Success i' -> (* We do have sk1[] = sk2[]: we now unify ?ev1 and ?ev2 *) (* Note that ?ev1 and ?ev2, may have been instantiated in the meantime *) - let ev1' = EConstr.of_constr (whd_evar i' (mkEvar ev1)) in - if EConstr.isEvar i' ev1' then - solve_simple_eqn (to_conv_fun (evar_conv_x ts)) env i' - (position_problem true pbty,EConstr.destEvar i' ev1', term2) + let ev1' = EConstr.of_constr (whd_evar i' (EConstr.Unsafe.to_constr (mkEvar ev1))) in + if isEvar i' ev1' then + solve_simple_eqn (evar_conv_x ts) env i' + (position_problem true pbty,destEvar i' ev1', term2) else evar_eqappr_x ts env evd pbty ((ev1', sk1), csts1) ((term2, sk2), csts2) | Some (r,[]), Success i' -> (* We have sk1'[] = sk2[] for some sk1' s.t. sk1[]=sk1'[r[]] *) (* we now unify r[?ev1] and ?ev2 *) - let ev2' = EConstr.of_constr (whd_evar i' (mkEvar ev2)) in - if EConstr.isEvar i' ev2' then - solve_simple_eqn (to_conv_fun (evar_conv_x ts)) env i' - (position_problem false pbty,EConstr.destEvar i' ev2',Stack.zip evd (term1,r)) + let ev2' = EConstr.of_constr (whd_evar i' (EConstr.Unsafe.to_constr (mkEvar ev2))) in + if isEvar i' ev2' then + solve_simple_eqn (evar_conv_x ts) env i' + (position_problem false pbty,destEvar i' ev2',Stack.zip evd (term1,r)) else evar_eqappr_x ts env evd pbty ((ev2', sk1), csts1) ((term2, sk2), csts2) @@ -552,10 +549,10 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty (* Symmetrically *) (* We have sk1[] = sk2'[] for some sk2' s.t. sk2[]=sk2'[r[]] *) (* we now unify ?ev1 and r[?ev2] *) - let ev1' = EConstr.of_constr (whd_evar i' (mkEvar ev1)) in - if EConstr.isEvar i' ev1' then - solve_simple_eqn (to_conv_fun (evar_conv_x ts)) env i' - (position_problem true pbty,EConstr.destEvar i' ev1',Stack.zip evd (term2,r)) + let ev1' = EConstr.of_constr (whd_evar i' (EConstr.Unsafe.to_constr (mkEvar ev1))) in + if isEvar i' ev1' then + solve_simple_eqn (evar_conv_x ts) env i' + (position_problem true pbty,destEvar i' ev1',Stack.zip evd (term2,r)) else evar_eqappr_x ts env evd pbty ((ev1', sk1), csts1) ((term2, sk2), csts2) | None, (UnifFailure _ as x) -> @@ -592,9 +589,9 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty if Evar.equal sp1 sp2 then match ise_stack2 false env i (evar_conv_x ts) sk1 sk2 with |None, Success i' -> - Success (solve_refl (to_conv_fun (fun env i pbty a1 a2 -> - is_success (evar_conv_x ts env i pbty a1 a2))) - env i' (position_problem true pbty) sp1 (Array.map EConstr.of_constr al1) (Array.map EConstr.of_constr al2)) + Success (solve_refl (fun env i pbty a1 a2 -> + is_success (evar_conv_x ts env i pbty a1 a2)) + env i' (position_problem true pbty) sp1 al1 al2) |_, (UnifFailure _ as x) -> x |Some _, _ -> UnifFailure (i,NotSameArgSize) else UnifFailure (i,NotSameHead) @@ -602,13 +599,13 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty ise_try evd [f1; f2] | Flexible ev1, MaybeFlexible v2 -> - flex_maybeflex true (fst ev1, Array.map EConstr.of_constr (snd ev1)) (appr1,csts1) (appr2,csts2) (EConstr.of_constr v2) + flex_maybeflex true ev1 (appr1,csts1) (appr2,csts2) v2 | MaybeFlexible v1, Flexible ev2 -> - flex_maybeflex false (fst ev2, Array.map EConstr.of_constr (snd ev2)) (appr2,csts2) (appr1,csts1) (EConstr.of_constr v1) + flex_maybeflex false ev2 (appr2,csts2) (appr1,csts1) v1 | MaybeFlexible v1, MaybeFlexible v2 -> begin - match kind_of_term (EConstr.Unsafe.to_constr term1), kind_of_term (EConstr.Unsafe.to_constr term2) with + match EConstr.kind evd term1, EConstr.kind evd term2 with | LetIn (na1,b1,t1,c'1), LetIn (na2,b2,t2,c'2) -> let f1 i = (* FO *) ise_and i @@ -617,14 +614,14 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty (fun i -> evar_conv_x ts env i CUMUL t2 t1)]); (fun i -> evar_conv_x ts env i CONV b1 b2); (fun i -> - let b = nf_evar i b1 in - let t = nf_evar i t1 in + let b = EConstr.to_constr i b1 in + let t = EConstr.to_constr i t1 in let na = Nameops.name_max na1 na2 in evar_conv_x ts (push_rel (RelDecl.LocalDef (na,b,t)) env) i pbty c'1 c'2); (fun i -> exact_ise_stack2 env i (evar_conv_x ts) sk1 sk2)] and f2 i = - let out1 = whd_betaiota_deltazeta_for_iota_state (fst ts) env i csts1 (EConstr.of_constr v1,sk1) - and out2 = whd_betaiota_deltazeta_for_iota_state (fst ts) env i csts2 (EConstr.of_constr v2,sk2) + let out1 = whd_betaiota_deltazeta_for_iota_state (fst ts) env i csts1 (v1,sk1) + and out2 = whd_betaiota_deltazeta_for_iota_state (fst ts) env i csts2 (v2,sk2) in evar_eqappr_x ts env i pbty out1 out2 in ise_try evd [f1; f2] @@ -636,8 +633,8 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty [(fun i -> evar_conv_x ts env i CONV c c'); (fun i -> exact_ise_stack2 env i (evar_conv_x ts) sk1 sk2)] and f2 i = - let out1 = whd_betaiota_deltazeta_for_iota_state (fst ts) env i csts1 (EConstr.of_constr v1,sk1) - and out2 = whd_betaiota_deltazeta_for_iota_state (fst ts) env i csts2 (EConstr.of_constr v2,sk2) + let out1 = whd_betaiota_deltazeta_for_iota_state (fst ts) env i csts1 (v1,sk1) + and out2 = whd_betaiota_deltazeta_for_iota_state (fst ts) env i csts2 (v2,sk2) in evar_eqappr_x ts env i pbty out1 out2 in ise_try evd [f1; f2] @@ -645,7 +642,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty (* Catch the p.c ~= p c' cases *) | Proj (p,c), Const (p',u) when eq_constant (Projection.constant p) p' -> let res = - try Some (EConstr.destApp evd (EConstr.of_constr (Retyping.expand_projection env evd p (EConstr.of_constr c) []))) + try Some (destApp evd (EConstr.of_constr (Retyping.expand_projection env evd p c []))) with Retyping.RetypeError _ -> None in (match res with @@ -656,7 +653,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty | Const (p,u), Proj (p',c') when eq_constant p (Projection.constant p') -> let res = - try Some (EConstr.destApp evd (EConstr.of_constr (Retyping.expand_projection env evd p' (EConstr.of_constr c') []))) + try Some (destApp evd (EConstr.of_constr (Retyping.expand_projection env evd p' c' []))) with Retyping.RetypeError _ -> None in (match res with @@ -710,7 +707,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty let applicative_stack = fst (Stack.strip_app sk2) in is_unnamed (fst (whd_betaiota_deltazeta_for_iota_state - (fst ts) env i Cst_stack.empty (EConstr.of_constr v2, applicative_stack))) in + (fst ts) env i Cst_stack.empty (v2, applicative_stack))) in let rhs_is_already_stuck = rhs_is_already_stuck || rhs_is_stuck_and_unnamed () in @@ -718,12 +715,12 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty && (not (Stack.not_purely_applicative sk1)) then evar_eqappr_x ~rhs_is_already_stuck ts env i pbty (whd_betaiota_deltazeta_for_iota_state - (fst ts) env i (Cst_stack.add_cst term1 csts1) (EConstr.of_constr v1,sk1)) + (fst ts) env i (Cst_stack.add_cst term1 csts1) (v1,sk1)) (appr2,csts2) else evar_eqappr_x ts env i pbty (appr1,csts1) (whd_betaiota_deltazeta_for_iota_state - (fst ts) env i (Cst_stack.add_cst term2 csts2) (EConstr.of_constr v2,sk2)) + (fst ts) env i (Cst_stack.add_cst term2 csts2) (v2,sk2)) in ise_try evd [f1; f2; f3] end @@ -731,14 +728,13 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty | Rigid, Rigid when EConstr.isLambda evd term1 && EConstr.isLambda evd term2 -> let (na1,c1,c'1) = EConstr.destLambda evd term1 in let (na2,c2,c'2) = EConstr.destLambda evd term2 in - let inj = EConstr.Unsafe.to_constr in assert app_empty; ise_and evd - [(fun i -> evar_conv_x ts env i CONV (inj c1) (inj c2)); + [(fun i -> evar_conv_x ts env i CONV c1 c2); (fun i -> - let c = nf_evar i (inj c1) in + let c = EConstr.to_constr i c1 in let na = Nameops.name_max na1 na2 in - evar_conv_x ts (push_rel (RelDecl.LocalAssum (na,c)) env) i CONV (inj c'1) (inj c'2))] + evar_conv_x ts (push_rel (RelDecl.LocalAssum (na,c)) env) i CONV c'1 c'2)] | Flexible ev1, Rigid -> flex_rigid true ev1 appr1 appr2 | Rigid, Flexible ev2 -> flex_rigid false ev2 appr2 appr1 @@ -752,7 +748,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty and f4 i = evar_eqappr_x ts env i pbty (whd_betaiota_deltazeta_for_iota_state - (fst ts) env i (Cst_stack.add_cst term1 csts1) (EConstr.of_constr v1,sk1)) + (fst ts) env i (Cst_stack.add_cst term1 csts1) (v1,sk1)) (appr2,csts2) in ise_try evd [f3; f4] @@ -766,19 +762,18 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty and f4 i = evar_eqappr_x ts env i pbty (appr1,csts1) (whd_betaiota_deltazeta_for_iota_state - (fst ts) env i (Cst_stack.add_cst term2 csts2) (EConstr.of_constr v2,sk2)) + (fst ts) env i (Cst_stack.add_cst term2 csts2) (v2,sk2)) in ise_try evd [f3; f4] (* Eta-expansion *) - | Rigid, _ when EConstr.isLambda evd term1 && (* if ever ill-typed: *) List.is_empty sk1 -> - eta env evd true sk1 (EConstr.Unsafe.to_constr term1) sk2 term2 + | Rigid, _ when isLambda evd term1 && (* if ever ill-typed: *) List.is_empty sk1 -> + eta env evd true sk1 term1 sk2 term2 - | _, Rigid when EConstr.isLambda evd term2 && (* if ever ill-typed: *) List.is_empty sk2 -> - eta env evd false sk2 (EConstr.Unsafe.to_constr term2) sk1 term1 + | _, Rigid when isLambda evd term2 && (* if ever ill-typed: *) List.is_empty sk2 -> + eta env evd false sk2 term2 sk1 term1 | Rigid, Rigid -> begin - let inj = EConstr.Unsafe.to_constr in match EConstr.kind evd term1, EConstr.kind evd term2 with | Sort s1, Sort s2 when app_empty -> @@ -794,11 +789,11 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty | Prod (n1,c1,c'1), Prod (n2,c2,c'2) when app_empty -> ise_and evd - [(fun i -> evar_conv_x ts env i CONV (inj c1) (inj c2)); + [(fun i -> evar_conv_x ts env i CONV c1 c2); (fun i -> - let c = nf_evar i (inj c1) in + let c = EConstr.to_constr i c1 in let na = Nameops.name_max n1 n2 in - evar_conv_x ts (push_rel (RelDecl.LocalAssum (na,c)) env) i pbty (inj c'1) (inj c'2))] + evar_conv_x ts (push_rel (RelDecl.LocalAssum (na,c)) env) i pbty c'1 c'2)] | Rel x1, Rel x2 -> if Int.equal x1 x2 then @@ -842,11 +837,10 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty else UnifFailure (evd,NotSameHead) | (Meta _, _) | (_, Meta _) -> - let inj = EConstr.Unsafe.to_constr in begin match ise_stack2 true env evd (evar_conv_x ts) sk1 sk2 with |_, (UnifFailure _ as x) -> x - |None, Success i' -> evar_conv_x ts env i' CONV (inj term1) (inj term2) - |Some (sk1',sk2'), Success i' -> evar_conv_x ts env i' CONV (inj (Stack.zip i' (term1,sk1'))) (inj (Stack.zip i' (term2,sk2'))) + |None, Success i' -> evar_conv_x ts env i' CONV term1 term2 + |Some (sk1',sk2'), Success i' -> evar_conv_x ts env i' CONV (Stack.zip i' (term1,sk1')) (Stack.zip i' (term2,sk2')) end | (Ind _ | Sort _ | Prod _ | CoFix _ | Fix _ | Rel _ | Var _ | Const _), _ -> @@ -884,38 +878,39 @@ and conv_record trs env evd (ctx,(h,h2),c,bs,(params,params1),(us,us2),(sk1,sk2) had to be initially resolved *) + let open EConstr in let evd = Evd.merge_context_set Evd.univ_flexible evd ctx in if Reductionops.Stack.compare_shape sk1 sk2 then let (evd',ks,_,test) = List.fold_left (fun (i,ks,m,test) b -> if match n with Some n -> Int.equal m n | None -> false then - let ty = Retyping.get_type_of env i (EConstr.of_constr t2) in - let test i = evar_conv_x trs env i CUMUL ty (substl ks b) in + let ty = EConstr.of_constr (Retyping.get_type_of env i t2) in + let test i = evar_conv_x trs env i CUMUL ty (Vars.substl ks b) in (i,t2::ks, m-1, test) else let dloc = (Loc.ghost,Evar_kinds.InternalHole) in let i = Sigma.Unsafe.of_evar_map i in - let Sigma (ev, i', _) = Evarutil.new_evar env i ~src:dloc (substl ks b) in + let Sigma (ev, i', _) = Evarutil.new_evar env i ~src:dloc (EConstr.Unsafe.to_constr (Vars.substl ks b)) in let i' = Sigma.to_evar_map i' in - (i', ev :: ks, m - 1,test)) + (i', EConstr.of_constr ev :: ks, m - 1,test)) (evd,[],List.length bs,fun i -> Success i) bs in let app = mkApp (c, Array.rev_of_list ks) in ise_and evd' [(fun i -> exact_ise_stack2 env i - (fun env' i' cpb x1 x -> evar_conv_x trs env' i' cpb x1 (substl ks x)) + (fun env' i' cpb x1 x -> evar_conv_x trs env' i' cpb x1 (Vars.substl ks x)) params1 params); (fun i -> exact_ise_stack2 env i - (fun env' i' cpb u1 u -> evar_conv_x trs env' i' cpb u1 (substl ks u)) + (fun env' i' cpb u1 u -> evar_conv_x trs env' i' cpb u1 (Vars.substl ks u)) us2 us); (fun i -> evar_conv_x trs env i CONV c1 app); (fun i -> exact_ise_stack2 env i (evar_conv_x trs) sk1 sk2); test; (fun i -> evar_conv_x trs env i CONV h2 - (fst (decompose_app_vect i (EConstr.of_constr (substl ks h)))))] + (EConstr.of_constr (fst (decompose_app_vect i (Vars.substl ks h)))))] else UnifFailure(evd,(*dummy*)NotSameHead) and eta_constructor ts env evd sk1 ((ind, i), u) sk2 term2 = @@ -956,66 +951,69 @@ let set_evar_conv f = Hook.set evar_conv_hook_set f (* We assume here |l1| <= |l2| *) let first_order_unification ts env evd (ev1,l1) (term2,l2) = + let open EConstr in let (deb2,rest2) = Array.chop (Array.length l2-Array.length l1) l2 in ise_and evd (* First compare extra args for better failure message *) - [(fun i -> ise_array2 i (fun i -> evar_conv_x ts env i CONV) (Array.map EConstr.of_constr rest2) (Array.map EConstr.of_constr l1)); + [(fun i -> ise_array2 i (fun i -> evar_conv_x ts env i CONV) rest2 l1); (fun i -> (* Then instantiate evar unless already done by unifying args *) let t2 = mkApp(term2,deb2) in if is_defined i (fst ev1) then evar_conv_x ts env i CONV t2 (mkEvar ev1) else - let ev1 = (fst ev1, Array.map EConstr.of_constr (snd ev1)) in - solve_simple_eqn ~choose:true (to_conv_fun (evar_conv_x ts)) env i (None,ev1, EConstr.of_constr t2))] + solve_simple_eqn ~choose:true (evar_conv_x ts) env i (None,ev1,t2))] let choose_less_dependent_instance evk evd term args = let evi = Evd.find_undefined evd evk in let subst = make_pure_subst evi args in - let subst' = List.filter (fun (id,c) -> Term.eq_constr c term) subst in + let subst' = List.filter (fun (id,c) -> EConstr.eq_constr evd c term) subst in match subst' with | [] -> None - | (id, _) :: _ -> Some (Evd.define evk (mkVar id) evd) + | (id, _) :: _ -> Some (Evd.define evk (Constr.mkVar id) evd) let apply_on_subterm env evdref f c t = + let open EConstr in let rec applyrec (env,(k,c) as acc) t = (* By using eq_constr, we make an approximation, for instance, we *) (* could also be interested in finding a term u convertible to t *) (* such that c occurs in u *) - if e_eq_constr_univs evdref c t then f k + if e_eq_constr_univs evdref (EConstr.Unsafe.to_constr c) (EConstr.Unsafe.to_constr t) then f k else - match kind_of_term t with - | Evar (evk,args) when Evd.is_undefined !evdref evk -> + match EConstr.kind !evdref t with + | Evar (evk,args) -> let ctx = evar_filtered_context (Evd.find_undefined !evdref evk) in let g decl a = if is_local_assum decl then applyrec acc a else a in mkEvar (evk, Array.of_list (List.map2 g ctx (Array.to_list args))) | _ -> - map_constr_with_binders_left_to_right - (fun d (env,(k,c)) -> (push_rel d env, (k+1,lift 1 c))) - applyrec acc t + let self acc c = EConstr.Unsafe.to_constr (applyrec acc (EConstr.of_constr c)) in + EConstr.of_constr (map_constr_with_binders_left_to_right + (fun d (env,(k,c)) -> (push_rel d env, (k+1,Vars.lift 1 c))) + self acc (EConstr.Unsafe.to_constr t)) in applyrec (env,(0,c)) t let filter_possible_projections evd c ty ctxt args = (* Since args in the types will be replaced by holes, we count the fv of args to have a well-typed filter; don't know how necessary - it is however to have a well-typed filter here *) - let fv1 = free_rels evd (EConstr.of_constr (mkApp (c,args))) (* Hack: locally untyped *) in - let fv2 = collect_vars evd (EConstr.of_constr (mkApp (c,args))) in + it is however to have a well-typed filter here *) + let open EConstr in + let fv1 = free_rels evd (mkApp (c,args)) (* Hack: locally untyped *) in + let fv2 = collect_vars evd (mkApp (c,args)) in let len = Array.length args in - let tyvars = collect_vars evd (EConstr.of_constr ty) in + let tyvars = collect_vars evd ty in List.map_i (fun i decl -> let () = assert (i < len) in let a = Array.unsafe_get args i in (match decl with | NamedDecl.LocalAssum _ -> false - | NamedDecl.LocalDef (_,c,_) -> not (isRel c || isVar c)) || + | NamedDecl.LocalDef (_,c,_) -> not (isRel evd (EConstr.of_constr c) || isVar evd (EConstr.of_constr c))) || a == c || (* Here we make an approximation, for instance, we could also be *) (* interested in finding a term u convertible to c such that a occurs *) (* in u *) - isRel a && Int.Set.mem (destRel a) fv1 || - isVar a && Id.Set.mem (destVar a) fv2 || + isRel evd a && Int.Set.mem (destRel evd a) fv1 || + isVar evd a && Id.Set.mem (destVar evd a) fv2 || Id.Set.mem (NamedDecl.get_id decl) tyvars) 0 ctxt @@ -1042,6 +1040,7 @@ let set_solve_evars f = solve_evars := f exception TypingFailed of evar_map let second_order_matching ts env_rhs evd (evk,args) argoccs rhs = + let open EConstr in try let evi = Evd.find_undefined evd evk in let env_evar = evar_filtered_env evi in @@ -1050,7 +1049,7 @@ let second_order_matching ts env_rhs evd (evk,args) argoccs rhs = let instance = List.map mkVar (List.map NamedDecl.get_id ctxt) in let rec make_subst = function - | decl'::ctxt', c::l, occs::occsl when isVarId (NamedDecl.get_id decl') c -> + | decl'::ctxt', c::l, occs::occsl when isVarId evd (NamedDecl.get_id decl') c -> begin match occs with | Some _ -> error "Cannot force abstraction on identity instance." @@ -1059,9 +1058,9 @@ let second_order_matching ts env_rhs evd (evk,args) argoccs rhs = end | decl'::ctxt', c::l, occs::occsl -> let id = NamedDecl.get_id decl' in - let t = NamedDecl.get_type decl' in + let t = EConstr.of_constr (NamedDecl.get_type decl') in let evs = ref [] in - let ty = Retyping.get_type_of env_rhs evd (EConstr.of_constr c) in + let ty = EConstr.of_constr (Retyping.get_type_of env_rhs evd c) in let filter' = filter_possible_projections evd c ty ctxt args in (id,t,c,ty,evs,Filter.make filter',occs) :: make_subst (ctxt',l,occsl) | _, _, [] -> [] @@ -1075,13 +1074,13 @@ let second_order_matching ts env_rhs evd (evk,args) argoccs rhs = | Some _ -> error "Selection of specific occurrences not supported" | None -> let evty = set_holes evdref cty subst in - let instance = Filter.filter_list filter instance in + let instance = List.map EConstr.Unsafe.to_constr (Filter.filter_list filter instance) in let evd = Sigma.Unsafe.of_evar_map !evdref in - let Sigma (ev, evd, _) = new_evar_instance sign evd evty ~filter instance in + let Sigma (ev, evd, _) = new_evar_instance sign evd (EConstr.Unsafe.to_constr evty) ~filter instance in let evd = Sigma.to_evar_map evd in evdref := evd; - evsref := (fst (destEvar ev),evty)::!evsref; - ev in + evsref := (fst (destEvar !evdref (EConstr.of_constr ev)),evty)::!evsref; + EConstr.of_constr ev in set_holes evdref (apply_on_subterm env_rhs evdref set_var c rhs) subst | [] -> rhs in @@ -1108,11 +1107,11 @@ let second_order_matching ts env_rhs evd (evk,args) argoccs rhs = (* We force abstraction over this unconstrained occurrence *) (* and we use typing to propagate this instantiation *) (* This is an arbitrary choice *) - let evd = Evd.define evk (mkVar id) evd in + let evd = Evd.define evk (Constr.mkVar id) evd in match evar_conv_x ts env_evar evd CUMUL idty evty with | UnifFailure _ -> error "Cannot find an instance" | Success evd -> - match reconsider_conv_pbs (to_conv_fun (evar_conv_x ts)) evd with + match reconsider_conv_pbs (evar_conv_x ts) evd with | UnifFailure _ -> error "Cannot find an instance" | Success evd -> evd @@ -1126,16 +1125,20 @@ let second_order_matching ts env_rhs evd (evk,args) argoccs rhs = force_instantiation evd !evsref | [] -> let evd = - try Evarsolve.check_evar_instance evd evk (EConstr.of_constr rhs) - (to_conv_fun (evar_conv_x full_transparent_state)) + try Evarsolve.check_evar_instance evd evk rhs + (evar_conv_x full_transparent_state) with IllTypedInstance _ -> raise (TypingFailed evd) in - Evd.define evk rhs evd + Evd.define evk (EConstr.Unsafe.to_constr rhs) evd in abstract_free_holes evd subst, true with TypingFailed evd -> evd, false +let to_pb (pb, env, t1, t2) = + (pb, env, EConstr.Unsafe.to_constr t1, EConstr.Unsafe.to_constr t2) + let second_order_matching_with_args ts env evd pbty ev l t = + let open EConstr in (* let evd,ev = evar_absorb_arguments env evd ev l in let argoccs = Array.map_to_list (fun _ -> None) (snd ev) in @@ -1144,18 +1147,19 @@ let second_order_matching_with_args ts env evd pbty ev l t = else UnifFailure (evd, ConversionFailed (env,mkApp(mkEvar ev,l),t)) if b then Success evd else *) - let pb = (pbty,env,mkApp(mkEvar ev,l),t) in + let pb = to_pb (pbty,env,mkApp(mkEvar ev,l),t) in UnifFailure (evd, CannotSolveConstraint (pb,ProblemBeyondCapabilities)) let apply_conversion_problem_heuristic ts env evd pbty t1 t2 = + let open EConstr in let t1 = apprec_nohdbeta ts env evd (whd_head_evar evd t1) in let t2 = apprec_nohdbeta ts env evd (whd_head_evar evd t2) in - let (term1,l1 as appr1) = try destApp t1 with DestKO -> (t1, [||]) in - let (term2,l2 as appr2) = try destApp t2 with DestKO -> (t2, [||]) in + let (term1,l1 as appr1) = try destApp evd t1 with DestKO -> (t1, [||]) in + let (term2,l2 as appr2) = try destApp evd t2 with DestKO -> (t2, [||]) in let app_empty = Array.is_empty l1 && Array.is_empty l2 in - match kind_of_term term1, kind_of_term term2 with + match EConstr.kind evd term1, EConstr.kind evd term2 with | Evar (evk1,args1), (Rel _|Var _) when app_empty - && List.for_all (fun a -> Term.eq_constr a term2 || isEvar a) + && List.for_all (fun a -> EConstr.eq_constr evd a term2 || isEvar evd a) (remove_instance_local_defs evd evk1 args1) -> (* The typical kind of constraint coming from pattern-matching return type inference *) @@ -1163,9 +1167,9 @@ let apply_conversion_problem_heuristic ts env evd pbty t1 t2 = | Some evd -> Success evd | None -> let reason = ProblemBeyondCapabilities in - UnifFailure (evd, CannotSolveConstraint ((pbty,env,t1,t2),reason))) + UnifFailure (evd, CannotSolveConstraint (to_pb (pbty,env,t1,t2),reason))) | (Rel _|Var _), Evar (evk2,args2) when app_empty - && List.for_all (fun a -> Term.eq_constr a term1 || isEvar a) + && List.for_all (fun a -> EConstr.eq_constr evd a term1 || isEvar evd a) (remove_instance_local_defs evd evk2 args2) -> (* The typical kind of constraint coming from pattern-matching return type inference *) @@ -1173,16 +1177,14 @@ let apply_conversion_problem_heuristic ts env evd pbty t1 t2 = | Some evd -> Success evd | None -> let reason = ProblemBeyondCapabilities in - UnifFailure (evd, CannotSolveConstraint ((pbty,env,t1,t2),reason))) + UnifFailure (evd, CannotSolveConstraint (to_pb (pbty,env,t1,t2),reason))) | Evar (evk1,args1), Evar (evk2,args2) when Evar.equal evk1 evk2 -> - let f env evd pbty x y = is_fconv ~reds:ts pbty env evd (EConstr.of_constr x) (EConstr.of_constr y) in - Success (solve_refl ~can_drop:true (to_conv_fun f) env evd - (position_problem true pbty) evk1 (Array.map EConstr.of_constr args1) (Array.map EConstr.of_constr args2)) + let f env evd pbty x y = is_fconv ~reds:ts pbty env evd x y in + Success (solve_refl ~can_drop:true f env evd + (position_problem true pbty) evk1 args1 args2) | Evar ev1, Evar ev2 when app_empty -> - let ev1 = (fst ev1, Array.map EConstr.of_constr (snd ev1)) in - let ev2 = (fst ev2, Array.map EConstr.of_constr (snd ev2)) in Success (solve_evar_evar ~force:true - (evar_define (to_conv_fun (evar_conv_x ts)) ~choose:true) (to_conv_fun (evar_conv_x ts)) env evd + (evar_define (evar_conv_x ts) ~choose:true) (evar_conv_x ts) env evd (position_problem true pbty) ev1 ev2) | Evar ev1,_ when Array.length l1 <= Array.length l2 -> (* On "?n t1 .. tn = u u1 .. u(n+p)", try first-order unification *) @@ -1244,9 +1246,9 @@ let rec solve_unconstrained_evars_with_candidates ts evd = | a::l -> try let conv_algo = evar_conv_x ts in - let evd = check_evar_instance evd evk (EConstr.of_constr a) (to_conv_fun conv_algo) in + let evd = check_evar_instance evd evk (EConstr.of_constr a) conv_algo in let evd = Evd.define evk a evd in - match reconsider_conv_pbs (to_conv_fun conv_algo) evd with + match reconsider_conv_pbs conv_algo evd with | Success evd -> solve_unconstrained_evars_with_candidates ts evd | UnifFailure _ -> aux l with @@ -1265,7 +1267,7 @@ let solve_unconstrained_impossible_cases env evd = let evd' = Evd.merge_context_set Evd.univ_flexible_alg ~loc evd' ctx in let ty = j_type j in let conv_algo = evar_conv_x full_transparent_state in - let evd' = check_evar_instance evd' evk (EConstr.of_constr ty) (to_conv_fun conv_algo) in + let evd' = check_evar_instance evd' evk (EConstr.of_constr ty) conv_algo in Evd.define evk ty evd' | _ -> evd') evd evd @@ -1275,7 +1277,7 @@ let consider_remaining_unif_problems env let rec aux evd pbs progress stuck = match pbs with | (pbty,env,t1,t2 as pb) :: pbs -> - (match apply_conversion_problem_heuristic ts env evd pbty t1 t2 with + (match apply_conversion_problem_heuristic ts env evd pbty (EConstr.of_constr t1) (EConstr.of_constr t2) with | Success evd' -> let (evd', rest) = extract_all_conv_pbs evd' in begin match rest with diff --git a/pretyping/evarconv.mli b/pretyping/evarconv.mli index 6f736e562..a0ff924ef 100644 --- a/pretyping/evarconv.mli +++ b/pretyping/evarconv.mli @@ -8,6 +8,7 @@ open Names open Term +open EConstr open Environ open Reductionops open Evd @@ -45,16 +46,16 @@ val check_problems_are_solved : env -> evar_map -> unit val check_conv_record : env -> evar_map -> state -> state -> Univ.universe_context_set * (constr * constr) - * constr * constr list * (EConstr.t Stack.t * EConstr.t Stack.t) * - (EConstr.t Stack.t * EConstr.t Stack.t) * - (EConstr.t Stack.t * EConstr.t Stack.t) * constr * + * constr * constr list * (constr Stack.t * constr Stack.t) * + (constr Stack.t * constr Stack.t) * + (constr Stack.t * constr Stack.t) * constr * (int option * constr) (** Try to solve problems of the form ?x[args] = c by second-order matching, using typing to select occurrences *) val second_order_matching : transparent_state -> env -> evar_map -> - existential -> occurrences option list -> constr -> evar_map * bool + EConstr.existential -> occurrences option list -> constr -> evar_map * bool (** Declare function to enforce evars resolution by using typing constraints *) diff --git a/pretyping/evardefine.ml b/pretyping/evardefine.ml index 8026ff3e4..f372dbf06 100644 --- a/pretyping/evardefine.ml +++ b/pretyping/evardefine.ml @@ -160,7 +160,7 @@ let define_evar_as_lambda env evd (evk,args) = evd, mkLambda (na, dom, evbody) let rec evar_absorb_arguments env evd (evk,args as ev) = function - | [] -> evd, (evk, Array.map EConstr.Unsafe.to_constr args) + | [] -> evd,ev | a::l -> let open EConstr in (* TODO: optimize and avoid introducing intermediate evars *) diff --git a/pretyping/evardefine.mli b/pretyping/evardefine.mli index f6d0efba6..f7bf4636b 100644 --- a/pretyping/evardefine.mli +++ b/pretyping/evardefine.mli @@ -27,7 +27,7 @@ val mk_valcon : EConstr.constr -> val_constraint [?y[vars1:=args1,vars:=args]] with [vars1 |- ?x:=\vars.?y[vars1:=vars1,vars:=vars]] *) val evar_absorb_arguments : env -> evar_map -> EConstr.existential -> EConstr.constr list -> - evar_map * existential + evar_map * EConstr.existential val split_tycon : Loc.t -> env -> evar_map -> type_constraint -> diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml index fbba682fc..570f95324 100644 --- a/pretyping/pretyping.ml +++ b/pretyping/pretyping.ml @@ -626,7 +626,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre let fixi = match fixkind with | GFix (vn,i) -> i | GCoFix i -> i - in e_conv env.ExtraEnv.env evdref ftys.(fixi) t + in e_conv env.ExtraEnv.env evdref (EConstr.of_constr ftys.(fixi)) (EConstr.of_constr t) | None -> true in (* Note: bodies are not used by push_rec_types, so [||] is safe *) @@ -732,7 +732,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre match candargs with | [] -> [], j_val hj | arg :: args -> - if e_conv env.ExtraEnv.env evdref (j_val hj) arg then + if e_conv env.ExtraEnv.env evdref (EConstr.of_constr (j_val hj)) (EConstr.of_constr arg) then args, nf_evar !evdref (j_val hj) else [], j_val hj in @@ -1088,7 +1088,7 @@ and pretype_type k0 resolve_tc valcon (env : ExtraEnv.t) evdref lvar = function match valcon with | None -> tj | Some v -> - if e_cumul env.ExtraEnv.env evdref v tj.utj_val then tj + if e_cumul env.ExtraEnv.env evdref (EConstr.of_constr v) (EConstr.of_constr tj.utj_val) then tj else error_unexpected_type ~loc:(loc_of_glob_constr c) env.ExtraEnv.env !evdref tj.utj_val v diff --git a/pretyping/typeclasses.ml b/pretyping/typeclasses.ml index 8c03329e2..11f71ee02 100644 --- a/pretyping/typeclasses.ml +++ b/pretyping/typeclasses.ml @@ -165,7 +165,7 @@ let rec is_class_type evd c = match kind_of_term c with | Prod (_, _, t) -> is_class_type evd t | Evar (e, _) when Evd.is_defined evd e -> - is_class_type evd (Evarutil.whd_head_evar evd c) + is_class_type evd (EConstr.Unsafe.to_constr (Evarutil.whd_head_evar evd (EConstr.of_constr c))) | _ -> is_class_constr c let is_class_evar evd evi = diff --git a/pretyping/typing.ml b/pretyping/typing.ml index 1dcb5f945..64264cf08 100644 --- a/pretyping/typing.ml +++ b/pretyping/typing.ml @@ -57,7 +57,7 @@ let e_judge_of_apply env evdref funj argjv = | hj::restjl -> match EConstr.kind !evdref (EConstr.of_constr (whd_all env !evdref typ)) with | Prod (_,c1,c2) -> - if Evarconv.e_cumul env evdref hj.uj_type (EConstr.Unsafe.to_constr c1) then + if Evarconv.e_cumul env evdref (EConstr.of_constr hj.uj_type) c1 then apply_rec (n+1) (Vars.subst1 (EConstr.of_constr hj.uj_val) c2) restjl else error_cant_apply_bad_type env (n, EConstr.Unsafe.to_constr c1, hj.uj_type) funj argjv @@ -75,7 +75,7 @@ let e_check_branch_types env evdref (ind,u) cj (lfj,explft) = if not (Int.equal (Array.length lfj) (Array.length explft)) then error_number_branches env cj (Array.length explft); for i = 0 to Array.length explft - 1 do - if not (Evarconv.e_cumul env evdref lfj.(i).uj_type explft.(i)) then + if not (Evarconv.e_cumul env evdref (EConstr.of_constr lfj.(i).uj_type) (EConstr.of_constr explft.(i))) then error_ill_formed_branch env cj.uj_val ((ind,i+1),u) lfj.(i).uj_type explft.(i) done @@ -91,7 +91,7 @@ let e_is_correct_arity env evdref c pj ind specif params = let pt' = whd_all env !evdref (EConstr.of_constr pt) in match kind_of_term pt', ar with | Prod (na1,a1,t), (LocalAssum (_,a1'))::ar' -> - if not (Evarconv.e_cumul env evdref a1 a1') then error (); + if not (Evarconv.e_cumul env evdref (EConstr.of_constr a1) (EConstr.of_constr a1')) then error (); srec (push_rel (LocalAssum (na1,a1)) env) t ar' | Sort s, [] -> if not (Sorts.List.mem (Sorts.family s) allowed_sorts) @@ -131,8 +131,8 @@ let check_type_fixpoint loc env evdref lna lar vdefj = let lt = Array.length vdefj in if Int.equal (Array.length lar) lt then for i = 0 to lt-1 do - if not (Evarconv.e_cumul env evdref (vdefj.(i)).uj_type - (lift lt lar.(i))) then + if not (Evarconv.e_cumul env evdref (EConstr.of_constr (vdefj.(i)).uj_type) + (EConstr.of_constr (lift lt lar.(i)))) then Pretype_errors.error_ill_typed_rec_body ~loc env !evdref i lna vdefj lar done @@ -150,7 +150,7 @@ let check_allowed_sort env sigma ind c p = let e_judge_of_cast env evdref cj k tj = let expected_type = tj.utj_val in - if not (Evarconv.e_cumul env evdref cj.uj_type expected_type) then + if not (Evarconv.e_cumul env evdref (EConstr.of_constr cj.uj_type) (EConstr.of_constr expected_type)) then error_actual_type env cj expected_type; { uj_val = mkCast (cj.uj_val, k, expected_type); uj_type = expected_type } @@ -282,7 +282,7 @@ and execute_array env evdref = Array.map (execute env evdref) let e_check env evdref c t = let env = enrich_env env evdref in let j = execute env evdref c in - if not (Evarconv.e_cumul env evdref j.uj_type t) then + if not (Evarconv.e_cumul env evdref (EConstr.of_constr j.uj_type) (EConstr.of_constr t)) then error_actual_type env j (nf_evar !evdref t) (* Type of a constr *) @@ -328,4 +328,4 @@ let e_solve_evars env evdref c = (* side-effect on evdref *) nf_evar !evdref c -let _ = Evarconv.set_solve_evars e_solve_evars +let _ = Evarconv.set_solve_evars (fun env evdref c -> EConstr.of_constr (e_solve_evars env evdref (EConstr.Unsafe.to_constr c))) diff --git a/pretyping/unification.ml b/pretyping/unification.ml index b8c9a93db..ac2f14051 100644 --- a/pretyping/unification.ml +++ b/pretyping/unification.ml @@ -123,7 +123,7 @@ let abstract_list_all_with_dependencies env evd typ c l = let argoccs = set_occurrences_of_last_arg (Array.sub (snd ev') 0 n) in let evd,b = Evarconv.second_order_matching empty_transparent_state - env evd ev' argoccs c in + env evd ev' argoccs (EConstr.of_constr c) in if b then let p = nf_evar evd ev in evd, p @@ -607,7 +607,7 @@ let check_compatibility env pbty flags (sigma,metasubst,evarsubst) tyM tyN = match subst_defined_metas_evars (metasubst,[]) tyN with | None -> sigma | Some n -> - if is_ground_term sigma m && is_ground_term sigma n then + if is_ground_term sigma (EConstr.of_constr m) && is_ground_term sigma (EConstr.of_constr n) then let sigma, b = infer_conv ~pb:pbty ~ts:flags.modulo_delta_types env sigma m n in if b then sigma else error_cannot_unify env sigma (m,n) @@ -659,8 +659,8 @@ let eta_constructor_app env f l1 term = let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb flags m n = let rec unirec_rec (curenv,nb as curenvnb) pb opt ((sigma,metasubst,evarsubst) as substn) curm curn = - let cM = Evarutil.whd_head_evar sigma curm - and cN = Evarutil.whd_head_evar sigma curn in + let cM = EConstr.Unsafe.to_constr (Evarutil.whd_head_evar sigma (EConstr.of_constr curm)) + and cN = EConstr.Unsafe.to_constr (Evarutil.whd_head_evar sigma (EConstr.of_constr curn)) in let () = if !debug_unification then Feedback.msg_debug (Termops.print_constr_env curenv cM ++ str" ~= " ++ Termops.print_constr_env curenv cN) @@ -964,7 +964,7 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb let sigma, b = infer_conv ~pb ~ts:convflags curenv sigma m1 n1 in if b then Some (sigma, metasubst, evarsubst) else - if is_ground_term sigma m1 && is_ground_term sigma n1 then + if is_ground_term sigma (EConstr.of_constr m1) && is_ground_term sigma (EConstr.of_constr n1) then error_cannot_unify curenv sigma (cM,cN) else None in @@ -1036,12 +1036,12 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb List.fold_left (fun (evd,ks,m) b -> if match n with Some n -> Int.equal m n | None -> false then - (evd,t2::ks, m-1) + (evd,EConstr.Unsafe.to_constr t2::ks, m-1) else let mv = new_meta () in let evd' = meta_declare mv (substl ks b) evd in (evd', mkMeta mv :: ks, m - 1)) - (sigma,[],List.length bs) bs + (sigma,[],List.length bs) (List.map EConstr.Unsafe.to_constr bs) in try let opt' = {opt with with_types = false} in @@ -1053,9 +1053,9 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb (fun s u1 u -> unirec_rec curenvnb pb opt' s (inj u1) (substl ks (inj u))) substn params1 params in let (substn,_,_) = Reductionops.Stack.fold2 (fun s u1 u2 -> unirec_rec curenvnb pb opt' s (inj u1) (inj u2)) substn ts ts1 in - let app = mkApp (c, Array.rev_of_list ks) in + let app = mkApp (EConstr.Unsafe.to_constr c, Array.rev_of_list ks) in (* let substn = unirec_rec curenvnb pb b false substn t cN in *) - unirec_rec curenvnb pb opt' substn c1 app + unirec_rec curenvnb pb opt' substn (EConstr.Unsafe.to_constr c1) app with Invalid_argument "Reductionops.Stack.fold2" -> error_cannot_unify (fst curenvnb) sigma (cM,cN) else error_cannot_unify (fst curenvnb) sigma (cM,cN) @@ -1271,11 +1271,8 @@ let order_metas metas = (* Solve an equation ?n[x1=u1..xn=un] = t where ?n is an evar *) -let to_conv_fun f = (); fun env sigma pb c1 c2 -> - f env sigma pb (EConstr.Unsafe.to_constr c1) (EConstr.Unsafe.to_constr c2) - let solve_simple_evar_eqn ts env evd ev rhs = - match solve_simple_eqn (to_conv_fun (Evarconv.evar_conv_x ts)) env evd (None,ev,EConstr.of_constr rhs) with + match solve_simple_eqn (Evarconv.evar_conv_x ts) env evd (None,ev,EConstr.of_constr rhs) with | UnifFailure (evd,reason) -> error_cannot_unify env evd ~reason (EConstr.Unsafe.to_constr (EConstr.mkEvar ev),rhs); | Success evd -> diff --git a/proofs/evar_refiner.ml b/proofs/evar_refiner.ml index d4b308bbe..af5fa921f 100644 --- a/proofs/evar_refiner.ml +++ b/proofs/evar_refiner.ml @@ -32,7 +32,7 @@ let define_and_solve_constraints evk c env evd = match List.fold_left (fun p (pbty,env,t1,t2) -> match p with - | Success evd -> Evarconv.evar_conv_x full_transparent_state env evd pbty t1 t2 + | Success evd -> Evarconv.evar_conv_x full_transparent_state env evd pbty (EConstr.of_constr t1) (EConstr.of_constr t2) | UnifFailure _ as x -> x) (Success evd) pbs with diff --git a/proofs/pfedit.ml b/proofs/pfedit.ml index c7f5efd5a..67e216745 100644 --- a/proofs/pfedit.ml +++ b/proofs/pfedit.ml @@ -232,7 +232,7 @@ let solve_by_implicit_tactic env sigma evk = let tac = Proofview.tclTHEN tac (Proofview.tclEXTEND [] (Proofview.tclZERO (CErrors.UserError (None,Pp.str"Proof is not complete."))) []) in (try let c = Evarutil.nf_evars_universes sigma evi.evar_concl in - if Evarutil.has_undefined_evars sigma c then raise Exit; + if Evarutil.has_undefined_evars sigma (EConstr.of_constr c) then raise Exit; let (ans, _, ctx) = build_by_tactic env (Evd.evar_universe_context sigma) c tac in let sigma = Evd.set_universe_context sigma ctx in diff --git a/stm/stm.ml b/stm/stm.ml index 23d68c4b8..6012e3d2d 100644 --- a/stm/stm.ml +++ b/stm/stm.ml @@ -1703,9 +1703,10 @@ end = struct (* {{{ *) Future.purify (fun () -> let _,_,_,_,sigma0 = Proof.proof (Proof_global.give_me_the_proof ()) in let g = Evd.find sigma0 r_goal in + let is_ground c = Evarutil.is_ground_term sigma0 (EConstr.of_constr c) in if not ( - Evarutil.is_ground_term sigma0 Evd.(evar_concl g) && - List.for_all (Context.Named.Declaration.for_all (Evarutil.is_ground_term sigma0)) + is_ground Evd.(evar_concl g) && + List.for_all (Context.Named.Declaration.for_all is_ground) Evd.(evar_context g)) then CErrors.user_err ~hdr:"STM" (strbrk("the par: goal selector supports ground "^ @@ -1719,7 +1720,7 @@ end = struct (* {{{ *) | Evd.Evar_empty -> RespNoProgress | Evd.Evar_defined t -> let t = Evarutil.nf_evar sigma t in - if Evarutil.is_ground_term sigma t then + if Evarutil.is_ground_term sigma (EConstr.of_constr t) then RespBuiltSubProof (t, Evd.evar_universe_context sigma) else CErrors.user_err ~hdr:"STM" (str"The solution is not ground") end) () diff --git a/tactics/class_tactics.ml b/tactics/class_tactics.ml index 6fb90e7af..a31e581e8 100644 --- a/tactics/class_tactics.ml +++ b/tactics/class_tactics.ml @@ -264,7 +264,7 @@ let unify_resolve_refine poly flags = let sigma' = let evdref = ref sigma' in if not (Evarconv.e_cumul env ~ts:flags.core_unify_flags.modulo_delta - evdref cl.cl_concl concl) then + evdref (EConstr.of_constr cl.cl_concl) (EConstr.of_constr concl)) then Type_errors.error_actual_type env {Environ.uj_val = term; Environ.uj_type = cl.cl_concl} concl; @@ -1506,7 +1506,7 @@ let not_evar c = | _ -> Proofview.tclUNIT () let is_ground c gl = - if Evarutil.is_ground_term (project gl) c then tclIDTAC gl + if Evarutil.is_ground_term (project gl) (EConstr.of_constr c) then tclIDTAC gl else tclFAIL 0 (str"Not ground") gl let autoapply c i gl = diff --git a/tactics/equality.ml b/tactics/equality.ml index e87746a28..17038e42d 100644 --- a/tactics/equality.ml +++ b/tactics/equality.ml @@ -626,7 +626,7 @@ let replace_using_leibniz clause c1 c2 l2r unsafe try_prove_eq_opt = let evd = if unsafe then Some (Tacmach.New.project gl) else - try Some (Evarconv.the_conv_x (Proofview.Goal.env gl) t1 t2 (Tacmach.New.project gl)) + try Some (Evarconv.the_conv_x (Proofview.Goal.env gl) (EConstr.of_constr t1) (EConstr.of_constr t2) (Tacmach.New.project gl)) with Evarconv.UnableToUnify _ -> None in match evd with @@ -1167,7 +1167,7 @@ let sig_clausal_form env sigma sort_of_ty siglen ty dflt = (* is the default value typable with the expected type *) let dflt_typ = unsafe_type_of env sigma dflt in try - let () = evdref := Evarconv.the_conv_x_leq env dflt_typ p_i !evdref in + let () = evdref := Evarconv.the_conv_x_leq env (EConstr.of_constr dflt_typ) (EConstr.of_constr p_i) !evdref in let () = evdref := Evarconv.consider_remaining_unif_problems env !evdref in dflt with Evarconv.UnableToUnify _ -> @@ -1185,7 +1185,7 @@ let sig_clausal_form env sigma sort_of_ty siglen ty dflt = with | Some w -> let w_type = unsafe_type_of env sigma w in - if Evarconv.e_cumul env evdref w_type a then + if Evarconv.e_cumul env evdref (EConstr.of_constr w_type) (EConstr.of_constr a) then let exist_term = Evarutil.evd_comb1 (Evd.fresh_global env) evdref sigdata.intro in applist(exist_term,[a;p_i_minus_1;w;tuple_tail]) else diff --git a/tactics/tactics.ml b/tactics/tactics.ml index 1c10cdfea..c2163a274 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -3731,7 +3731,7 @@ let specialize_eqs id gl = let ty = Tacmach.pf_get_hyp_typ gl id in let evars = ref (project gl) in let unif env evars c1 c2 = - compare_upto_variables c1 c2 && Evarconv.e_conv env evars c1 c2 + compare_upto_variables c1 c2 && Evarconv.e_conv env evars (EConstr.of_constr c1) (EConstr.of_constr c2) in let rec aux in_eqs ctx acc ty = match kind_of_term ty with @@ -4275,7 +4275,7 @@ let check_expected_type env sigma (elimc,bl) elimt = let sigma,cl = make_evar_clause env sigma ~len:(n - 1) elimt in let sigma = solve_evar_clause env sigma true cl bl in let (_,u,_) = destProd cl.cl_concl in - fun t -> Evarconv.e_cumul env (ref sigma) t u + fun t -> Evarconv.e_cumul env (ref sigma) (EConstr.of_constr t) (EConstr.of_constr u) let check_enough_applied env sigma elim = let sigma = Sigma.to_evar_map sigma in diff --git a/toplevel/vernacentries.ml b/toplevel/vernacentries.ml index c39b4dc25..33b96bc37 100644 --- a/toplevel/vernacentries.ml +++ b/toplevel/vernacentries.ml @@ -458,7 +458,7 @@ let start_proof_and_print k l hook = let env = Evd.evar_filtered_env evi in try let concl = Evarutil.nf_evars_universes sigma evi.Evd.evar_concl in - if Evarutil.has_undefined_evars sigma concl then raise Exit; + if Evarutil.has_undefined_evars sigma (EConstr.of_constr concl) then raise Exit; let c, _, ctx = Pfedit.build_by_tactic env (Evd.evar_universe_context sigma) concl (Tacticals.New.tclCOMPLETE tac) @@ -1578,7 +1578,7 @@ let vernac_check_may_eval redexp glopt rc = let env = Environ.push_context uctx (Evarutil.nf_env_evar sigma' env) in let c = nf c in let j = - if Evarutil.has_undefined_evars sigma' c then + if Evarutil.has_undefined_evars sigma' (EConstr.of_constr c) then Evarutil.j_nf_evar sigma' (Retyping.get_judgment_of env sigma' (EConstr.of_constr c)) else (* OK to call kernel which does not support evars *) -- cgit v1.2.3 From e27949240f5b1ee212e7d0fe3326a21a13c4abb0 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sun, 6 Nov 2016 17:21:44 +0100 Subject: Typing API using EConstr. --- ltac/evar_tactics.ml | 2 +- ltac/rewrite.ml | 18 +++--- ltac/tacinterp.ml | 4 +- plugins/cc/cctac.ml | 10 ++-- plugins/decl_mode/decl_proof_instr.ml | 4 +- plugins/firstorder/instances.ml | 2 +- plugins/funind/functional_principles_proofs.ml | 10 ++-- plugins/funind/functional_principles_types.ml | 10 ++-- plugins/funind/glob_term_to_relation.ml | 18 +++--- plugins/funind/indfun.ml | 2 +- plugins/funind/invfun.ml | 8 +-- plugins/funind/recdef.ml | 2 +- plugins/setoid_ring/newring.ml | 8 +-- pretyping/cases.ml | 10 ++-- pretyping/coercion.ml | 12 ++-- pretyping/pretyping.ml | 10 ++-- pretyping/tacred.ml | 2 +- pretyping/typing.ml | 80 ++++++++++++++++---------- pretyping/typing.mli | 16 +++--- pretyping/unification.ml | 4 +- proofs/clenv.ml | 2 +- proofs/logic.ml | 4 +- proofs/refine.ml | 8 +-- proofs/tacmach.ml | 8 +-- tactics/autorewrite.ml | 4 +- tactics/equality.ml | 16 +++--- tactics/hints.ml | 2 +- tactics/inv.ml | 4 +- tactics/tactics.ml | 26 ++++----- toplevel/class.ml | 2 +- toplevel/command.ml | 8 +-- 31 files changed, 169 insertions(+), 147 deletions(-) diff --git a/ltac/evar_tactics.ml b/ltac/evar_tactics.ml index c5b26e6d5..5d3b2b886 100644 --- a/ltac/evar_tactics.ml +++ b/ltac/evar_tactics.ml @@ -77,7 +77,7 @@ let let_evar name typ = let sigma = Tacmach.New.project gl in let env = Proofview.Goal.env gl in let sigma = ref sigma in - let _ = Typing.e_sort_of env sigma typ in + let _ = Typing.e_sort_of env sigma (EConstr.of_constr typ) in let sigma = Sigma.Unsafe.of_evar_map !sigma in let id = match name with | Names.Anonymous -> diff --git a/ltac/rewrite.ml b/ltac/rewrite.ml index 076e4c05e..5703687c4 100644 --- a/ltac/rewrite.ml +++ b/ltac/rewrite.ml @@ -113,7 +113,7 @@ let extends_undefined evars evars' = let app_poly_check env evars f args = let (evars, cstrs), fc = f evars in let evdref = ref evars in - let t = Typing.e_solve_evars env evdref (mkApp (fc, args)) in + let t = Typing.e_solve_evars env evdref (EConstr.of_constr (mkApp (fc, args))) in (!evdref, cstrs), t let app_poly_nocheck env evars f args = @@ -382,7 +382,7 @@ end let type_app_poly env env evd f args = let evars, c = app_poly_nocheck env evd f args in - let evd', t = Typing.type_of env (goalevars evars) c in + let evd', t = Typing.type_of env (goalevars evars) (EConstr.of_constr c) in (evd', cstrevars evars), c module PropGlobal = struct @@ -485,7 +485,7 @@ let rec decompose_app_rel env evd t = | App (f, [||]) -> assert false | App (f, [|arg|]) -> let (f', argl, argr) = decompose_app_rel env evd arg in - let ty = Typing.unsafe_type_of env evd argl in + let ty = Typing.unsafe_type_of env evd (EConstr.of_constr argl) in let f'' = mkLambda (Name default_dependent_ident, ty, mkLambda (Name (Id.of_string "y"), lift 1 ty, mkApp (lift 2 f, [| mkApp (lift 2 f', [| mkRel 2; mkRel 1 |]) |]))) @@ -787,7 +787,7 @@ let resolve_morphism env avoid oldt m ?(fnewt=fun x -> x) args args' (b,cstr) ev let morphargs, morphobjs = Array.chop first args in let morphargs', morphobjs' = Array.chop first args' in let appm = mkApp(m, morphargs) in - let appmtype = Typing.unsafe_type_of env (goalevars evars) appm in + let appmtype = Typing.unsafe_type_of env (goalevars evars) (EConstr.of_constr appm) in let cstrs = List.map (Option.map (fun r -> r.rew_car, get_opt_rew_rel r.rew_prf)) (Array.to_list morphobjs') @@ -1477,7 +1477,7 @@ type result = (evar_map * constr option * types) option option let cl_rewrite_clause_aux ?(abs=None) strat env avoid sigma concl is_hyp : result = let evdref = ref sigma in - let sort = Typing.e_sort_of env evdref concl in + let sort = Typing.e_sort_of env evdref (EConstr.of_constr concl) in let evars = (!evdref, Evar.Set.empty) in let evars, cstr = let prop, (evars, arrow) = @@ -1870,7 +1870,7 @@ let declare_projection n instance_id r = let sigma,c = Evd.fresh_global env sigma r in let ty = Retyping.get_type_of env sigma (EConstr.of_constr c) in let term = proper_projection c ty in - let sigma, typ = Typing.type_of env sigma term in + let sigma, typ = Typing.type_of env sigma (EConstr.of_constr term) in let ctx, typ = decompose_prod_assum typ in let typ = let n = @@ -1902,7 +1902,7 @@ let declare_projection n instance_id r = let build_morphism_signature env sigma m = let m,ctx = Constrintern.interp_constr env sigma m in let sigma = Evd.from_ctx ctx in - let t = Typing.unsafe_type_of env sigma m in + let t = Typing.unsafe_type_of env sigma (EConstr.of_constr m) in let cstrs = let rec aux t = match kind_of_term t with @@ -1932,7 +1932,7 @@ let build_morphism_signature env sigma m = let default_morphism sign m = let env = Global.env () in let sigma = Evd.from_env env in - let t = Typing.unsafe_type_of env sigma m in + let t = Typing.unsafe_type_of env sigma (EConstr.of_constr m) in let evars, _, sign, cstrs = PropGlobal.build_signature (sigma, Evar.Set.empty) env t (fst sign) (snd sign) in @@ -2126,7 +2126,7 @@ let setoid_proof ty fn fallback = begin try let rel, _, _ = decompose_app_rel env sigma concl in - let (sigma, t) = Typing.type_of env sigma rel in + let (sigma, t) = Typing.type_of env sigma (EConstr.of_constr rel) in let car = RelDecl.get_type (List.hd (fst (Reduction.dest_prod env t))) in (try init_relation_classes () with _ -> raise Not_found); fn env sigma car rel diff --git a/ltac/tacinterp.ml b/ltac/tacinterp.ml index 34faa028f..8f5ac7e76 100644 --- a/ltac/tacinterp.ml +++ b/ltac/tacinterp.ml @@ -791,7 +791,7 @@ let interp_may_eval f ist env sigma = function let ctxt = coerce_to_constr_context (Id.Map.find s ist.lfun) in let evdref = ref sigma in let c = subst_meta [Constr_matching.special_meta,ic] ctxt in - let c = Typing.e_solve_evars env evdref c in + let c = Typing.e_solve_evars env evdref (EConstr.of_constr c) in !evdref , c with | Not_found -> @@ -799,7 +799,7 @@ let interp_may_eval f ist env sigma = function (str "Unbound context identifier" ++ pr_id s ++ str".")) | ConstrTypeOf c -> let (sigma,c_interp) = f ist env sigma c in - Typing.type_of ~refresh:true env sigma c_interp + Typing.type_of ~refresh:true env sigma (EConstr.of_constr c_interp) | ConstrTerm c -> try f ist env sigma c diff --git a/plugins/cc/cctac.ml b/plugins/cc/cctac.ml index 36a96fdb5..58454eedf 100644 --- a/plugins/cc/cctac.ml +++ b/plugins/cc/cctac.ml @@ -50,7 +50,7 @@ let whd_delta env= (* decompose member of equality in an applicative format *) (** FIXME: evar leak *) -let sf_of env sigma c = e_sort_of env (ref sigma) c +let sf_of env sigma c = e_sort_of env (ref sigma) (EConstr.of_constr c) let rec decompose_term env sigma t= match kind_of_term (whd env t) with @@ -247,7 +247,7 @@ let new_refine c = Proofview.V82.tactic (refine c) let assert_before n c = Proofview.Goal.enter { enter = begin fun gl -> - let evm, _ = Tacmach.New.pf_apply type_of gl c in + let evm, _ = Tacmach.New.pf_apply type_of gl (EConstr.of_constr c) in Tacticals.New.tclTHEN (Proofview.V82.tactic (Refiner.tclEVARS evm)) (assert_before n c) end } @@ -340,7 +340,7 @@ let refute_tac c t1 t2 p = end } let refine_exact_check c gl = - let evm, _ = pf_apply type_of gl c in + let evm, _ = pf_apply type_of gl (EConstr.of_constr c) in Tacticals.tclTHEN (Refiner.tclEVARS evm) (Proofview.V82.of_tactic (exact_check c)) gl let convert_to_goal_tac c t1 t2 p = @@ -480,10 +480,10 @@ let mk_eq f c1 c2 k = Tacticals.New.pf_constr_of_global (Lazy.force f) (fun fc -> Proofview.Goal.enter { enter = begin fun gl -> let open Tacmach.New in - let evm, ty = pf_apply type_of gl c1 in + let evm, ty = pf_apply type_of gl (EConstr.of_constr c1) in let evm, ty = Evarsolve.refresh_universes (Some false) (pf_env gl) evm (EConstr.of_constr ty) in let term = mkApp (fc, [| ty; c1; c2 |]) in - let evm, _ = type_of (pf_env gl) evm term in + let evm, _ = type_of (pf_env gl) evm (EConstr.of_constr term) in Tacticals.New.tclTHEN (Proofview.V82.tactic (Refiner.tclEVARS evm)) (k term) end }) diff --git a/plugins/decl_mode/decl_proof_instr.ml b/plugins/decl_mode/decl_proof_instr.ml index c17c8dbb8..dcebf7806 100644 --- a/plugins/decl_mode/decl_proof_instr.ml +++ b/plugins/decl_mode/decl_proof_instr.ml @@ -429,7 +429,7 @@ let concl_refiner metas body gls = let concl = pf_concl gls in let evd = sig_sig gls in let env = pf_env gls in - let sort = family_of_sort (Typing.e_sort_of env (ref evd) concl) in + let sort = family_of_sort (Typing.e_sort_of env (ref evd) (EConstr.of_constr concl)) in let rec aux env avoid subst = function [] -> anomaly ~label:"concl_refiner" (Pp.str "cannot happen") | (n,typ)::rest -> @@ -437,7 +437,7 @@ let concl_refiner metas body gls = let x = id_of_name_using_hdchar env _A Anonymous in let _x = fresh_id avoid x gls in let nenv = Environ.push_named (LocalAssum (_x,_A)) env in - let asort = family_of_sort (Typing.e_sort_of nenv (ref evd) _A) in + let asort = family_of_sort (Typing.e_sort_of nenv (ref evd) (EConstr.of_constr _A)) in let nsubst = (n,mkVar _x)::subst in if List.is_empty rest then asort,_A,mkNamedLambda _x _A (subst_meta nsubst body) diff --git a/plugins/firstorder/instances.ml b/plugins/firstorder/instances.ml index a3513692c..44bdb585a 100644 --- a/plugins/firstorder/instances.ml +++ b/plugins/firstorder/instances.ml @@ -155,7 +155,7 @@ let left_instance_tac (inst,id) continue seq= it_mkLambda_or_LetIn (mkApp(idc,[|ot|])) rc in let evmap, _ = - try Typing.type_of (pf_env gl) evmap gt + try Typing.type_of (pf_env gl) evmap (EConstr.of_constr gt) with e when CErrors.noncritical e -> error "Untypable instance, maybe higher-order non-prenex quantification" in tclTHEN (Refiner.tclEVARS evmap) (Proofview.V82.of_tactic (generalize [gt])) gl) diff --git a/plugins/funind/functional_principles_proofs.ml b/plugins/funind/functional_principles_proofs.ml index 340dd2c28..0a7938069 100644 --- a/plugins/funind/functional_principles_proofs.ml +++ b/plugins/funind/functional_principles_proofs.ml @@ -329,7 +329,7 @@ let change_eq env sigma hyp_id (context:Context.Rel.t) x t end_of_type = let all_ids = pf_ids_of_hyps g in let new_ids,_ = list_chop ctxt_size all_ids in let to_refine = applist(witness_fun,List.rev_map mkVar new_ids) in - let evm, _ = pf_apply Typing.type_of g to_refine in + let evm, _ = pf_apply Typing.type_of g (EConstr.of_constr to_refine) in tclTHEN (Refiner.tclEVARS evm) (refine to_refine) g ) in @@ -544,7 +544,7 @@ let clean_hyp_with_heq ptes_infos eq_hyps hyp_id env sigma = tclIDTAC in try - scan_type [] (Typing.unsafe_type_of env sigma (mkVar hyp_id)), [hyp_id] + scan_type [] (Typing.unsafe_type_of env sigma (EConstr.mkVar hyp_id)), [hyp_id] with TOREMOVE -> thin [hyp_id],[] @@ -639,7 +639,7 @@ let instanciate_hyps_with_args (do_prove:Id.t list -> tactic) hyps args_id = fun g -> let prov_hid = pf_get_new_id hid g in let c = mkApp(mkVar hid,args) in - let evm, _ = pf_apply Typing.type_of g c in + let evm, _ = pf_apply Typing.type_of g (EConstr.of_constr c) in tclTHENLIST[ Refiner.tclEVARS evm; Proofview.V82.of_tactic (pose_proof (Name prov_hid) c); @@ -968,7 +968,7 @@ let generate_equation_lemma evd fnames f fun_num nb_params nb_args rec_args_num let eq_rhs = nf_betaiotazeta (mkApp(compose_lam params f_body_with_params_and_other_fun,Array.init (nb_params + nb_args) (fun i -> mkRel(nb_params + nb_args - i)))) in (* observe (str "eq_rhs " ++ pr_lconstr eq_rhs); *) let (type_ctxt,type_of_f),evd = - let evd,t = Typing.type_of ~refresh:true (Global.env ()) evd f + let evd,t = Typing.type_of ~refresh:true (Global.env ()) evd (EConstr.of_constr f) in decompose_prod_n_assum (nb_params + nb_args) t,evd @@ -1039,7 +1039,7 @@ let do_replace (evd:Evd.evar_map ref) params rec_arg_num rev_args_id f fun_num a (Constrintern.locate_reference (qualid_of_ident equation_lemma_id)) in evd:=evd'; - let _ = Typing.e_type_of ~refresh:true (Global.env ()) evd res in + let _ = Typing.e_type_of ~refresh:true (Global.env ()) evd (EConstr.of_constr res) in res in let nb_intro_to_do = nb_prod (project g) (EConstr.of_constr (pf_concl g)) in diff --git a/plugins/funind/functional_principles_types.ml b/plugins/funind/functional_principles_types.ml index 9637632a6..4b47b83af 100644 --- a/plugins/funind/functional_principles_types.ml +++ b/plugins/funind/functional_principles_types.ml @@ -283,7 +283,7 @@ let build_functional_principle (evd:Evd.evar_map ref) interactive_proof old_prin let new_princ_name = next_ident_away_in_goal (Id.of_string "___________princ_________") [] in - let _ = Typing.e_type_of ~refresh:true (Global.env ()) evd new_principle_type in + let _ = Typing.e_type_of ~refresh:true (Global.env ()) evd (EConstr.of_constr new_principle_type) in let hook = Lemmas.mk_hook (hook new_principle_type) in begin Lemmas.start_proof @@ -337,7 +337,7 @@ let generate_functional_principle (evd: Evd.evar_map ref) let evd',s = Evd.fresh_sort_in_family env evd' fam_sort in let name = Indrec.make_elimination_ident base_new_princ_name fam_sort in let evd',value = change_property_sort evd' s new_principle_type new_princ_name in - let evd' = fst (Typing.type_of ~refresh:true (Global.env ()) evd' value) in + let evd' = fst (Typing.type_of ~refresh:true (Global.env ()) evd' (EConstr.of_constr value)) in (* Pp.msgnl (str "new principle := " ++ pr_lconstr value); *) let ce = Declare.definition_entry ~poly:(Flags.is_universe_polymorphism ()) ~univs:(snd (Evd.universe_context evd')) value in ignore( @@ -488,7 +488,7 @@ let make_scheme evd (fas : (pconstant*glob_sort) list) : Safe_typing.private_con in let _ = evd := sigma in let l_schemes = - List.map (Typing.unsafe_type_of env sigma) schemes + List.map (EConstr.of_constr %> Typing.unsafe_type_of env sigma) schemes in let i = ref (-1) in let sorts = @@ -616,7 +616,7 @@ let build_scheme fas = in let evd',f = Evd.fresh_global (Global.env ()) !evd f_as_constant in let _ = evd := evd' in - let _ = Typing.e_type_of ~refresh:true (Global.env ()) evd f in + let _ = Typing.e_type_of ~refresh:true (Global.env ()) evd (EConstr.of_constr f) in (destConst f,sort) ) fas @@ -666,7 +666,7 @@ let build_case_scheme fa = Indrec.build_case_analysis_scheme_default env sigma ind sf in let sigma = Sigma.to_evar_map sigma in - let scheme_type = (Typing.unsafe_type_of env sigma ) scheme in + let scheme_type = (Typing.unsafe_type_of env sigma) (EConstr.of_constr scheme) in let sorts = (fun (_,_,x) -> Universes.new_sort_in_family (Pretyping.interp_elimination_sort x) diff --git a/plugins/funind/glob_term_to_relation.ml b/plugins/funind/glob_term_to_relation.ml index 92de4d873..38cd21684 100644 --- a/plugins/funind/glob_term_to_relation.ml +++ b/plugins/funind/glob_term_to_relation.ml @@ -503,7 +503,7 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return = The "value" of this branch is then simply [res] *) let rt_as_constr,ctx = Pretyping.understand env (Evd.from_env env) rt in - let rt_typ = Typing.unsafe_type_of env (Evd.from_env env) rt_as_constr in + let rt_typ = Typing.unsafe_type_of env (Evd.from_env env) (EConstr.of_constr rt_as_constr) in let res_raw_type = Detyping.detype false [] env (Evd.from_env env) rt_typ in let res = fresh_id args_res.to_avoid "_res" in let new_avoid = res::args_res.to_avoid in @@ -611,7 +611,7 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return = *) let v_res = build_entry_lc env funnames avoid v in let v_as_constr,ctx = Pretyping.understand env (Evd.from_env env) v in - let v_type = Typing.unsafe_type_of env (Evd.from_env env) v_as_constr in + let v_type = Typing.unsafe_type_of env (Evd.from_env env) (EConstr.of_constr v_as_constr) in let new_env = match n with Anonymous -> env @@ -627,7 +627,7 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return = build_entry_lc_from_case env funnames make_discr el brl avoid | GIf(_,b,(na,e_option),lhs,rhs) -> let b_as_constr,ctx = Pretyping.understand env (Evd.from_env env) b in - let b_typ = Typing.unsafe_type_of env (Evd.from_env env) b_as_constr in + let b_typ = Typing.unsafe_type_of env (Evd.from_env env) (EConstr.of_constr b_as_constr) in let (ind,_) = try Inductiveops.find_inductive env (Evd.from_env env) (EConstr.of_constr b_typ) with Not_found -> @@ -659,7 +659,7 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return = nal in let b_as_constr,ctx = Pretyping.understand env (Evd.from_env env) b in - let b_typ = Typing.unsafe_type_of env (Evd.from_env env) b_as_constr in + let b_typ = Typing.unsafe_type_of env (Evd.from_env env) (EConstr.of_constr b_as_constr) in let (ind,_) = try Inductiveops.find_inductive env (Evd.from_env env) (EConstr.of_constr b_typ) with Not_found -> @@ -706,7 +706,7 @@ and build_entry_lc_from_case env funname make_discr let types = List.map (fun (case_arg,_) -> let case_arg_as_constr,ctx = Pretyping.understand env (Evd.from_env env) case_arg in - Typing.unsafe_type_of env (Evd.from_env env) case_arg_as_constr + Typing.unsafe_type_of env (Evd.from_env env) (EConstr.of_constr case_arg_as_constr) ) el in (****** The next works only if the match is not dependent ****) @@ -753,7 +753,7 @@ and build_entry_lc_from_case_term env types funname make_discr patterns_to_preve List.fold_right (fun id acc -> let typ_of_id = - Typing.unsafe_type_of env_with_pat_ids (Evd.from_env env) (mkVar id) + Typing.unsafe_type_of env_with_pat_ids (Evd.from_env env) (EConstr.mkVar id) in let raw_typ_of_id = Detyping.detype false [] @@ -807,7 +807,7 @@ and build_entry_lc_from_case_term env types funname make_discr patterns_to_preve (fun id acc -> if Id.Set.mem id this_pat_ids then (Prod (Name id), - let typ_of_id = Typing.unsafe_type_of new_env (Evd.from_env env) (mkVar id) in + let typ_of_id = Typing.unsafe_type_of new_env (Evd.from_env env) (EConstr.mkVar id) in let raw_typ_of_id = Detyping.detype false [] new_env (Evd.from_env env) typ_of_id in @@ -1121,7 +1121,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt = let evd = (Evd.from_env env) in let t',ctx = Pretyping.understand env evd t in let evd = Evd.from_ctx ctx in - let type_t' = Typing.unsafe_type_of env evd t' in + let type_t' = Typing.unsafe_type_of env evd (EConstr.of_constr t') in let new_env = Environ.push_rel (LocalDef (n,t',type_t')) env in let new_b,id_to_exclude = rebuild_cons new_env @@ -1272,7 +1272,7 @@ let do_build_inductive let evd,env = Array.fold_right2 (fun id c (evd,env) -> - let evd,t = Typing.type_of env evd (mkConstU c) in + let evd,t = Typing.type_of env evd (EConstr.mkConstU c) in evd, Environ.push_named (LocalAssum (id,t)) (* try *) diff --git a/plugins/funind/indfun.ml b/plugins/funind/indfun.ml index a264c37c5..0743fc5d9 100644 --- a/plugins/funind/indfun.ml +++ b/plugins/funind/indfun.ml @@ -369,7 +369,7 @@ let generate_principle (evd:Evd.evar_map ref) pconstants on_error let evd = ref (Evd.from_env env) in let evd',uprinc = Evd.fresh_global env !evd princ in let _ = evd := evd' in - let princ_type = Typing.e_type_of ~refresh:true env evd uprinc in + let princ_type = Typing.e_type_of ~refresh:true env evd (EConstr.of_constr uprinc) in Functional_principles_types.generate_functional_principle evd interactive_proof diff --git a/plugins/funind/invfun.ml b/plugins/funind/invfun.ml index 9abe60402..e5286fb1f 100644 --- a/plugins/funind/invfun.ml +++ b/plugins/funind/invfun.ml @@ -134,7 +134,7 @@ let generate_type evd g_to_f f graph i = Evd.fresh_global (Global.env ()) !evd (Globnames.IndRef (fst (destInd graph))) in evd:=evd'; - let graph_arity = Typing.e_type_of (Global.env ()) evd graph in + let graph_arity = Typing.e_type_of (Global.env ()) evd (EConstr.of_constr graph) in let ctxt,_ = decompose_prod_assum graph_arity in let fun_ctxt,res_type = match ctxt with @@ -202,7 +202,7 @@ let find_induction_principle evd f = | None -> raise Not_found | Some rect_lemma -> let evd',rect_lemma = Evd.fresh_global (Global.env ()) !evd (Globnames.ConstRef rect_lemma) in - let evd',typ = Typing.type_of ~refresh:true (Global.env ()) evd' rect_lemma in + let evd',typ = Typing.type_of ~refresh:true (Global.env ()) evd' (EConstr.of_constr rect_lemma) in evd:=evd'; rect_lemma,typ @@ -449,7 +449,7 @@ let prove_fun_correct evd functional_induction funs_constr graphs_constr schemes "functional_induction" ( (fun gl -> let term = mkApp (mkVar principle_id,Array.of_list bindings) in - let gl', _ty = pf_eapply (Typing.type_of ~refresh:true) gl term in + let gl', _ty = pf_eapply (Typing.type_of ~refresh:true) gl (EConstr.of_constr term) in Proofview.V82.of_tactic (apply term) gl') )) (fun i g -> observe_tac ("proving branche "^string_of_int i) (prove_branche i) g ) @@ -792,7 +792,7 @@ let derive_correctness make_scheme functional_induction (funs: pconstant list) ( let type_info = (type_of_lemma_ctxt,type_of_lemma_concl) in graphs_constr.(i) <- graph; let type_of_lemma = Termops.it_mkProd_or_LetIn type_of_lemma_concl type_of_lemma_ctxt in - let _ = Typing.e_type_of (Global.env ()) evd type_of_lemma in + let _ = Typing.e_type_of (Global.env ()) evd (EConstr.of_constr type_of_lemma) in let type_of_lemma = nf_zeta (EConstr.of_constr type_of_lemma) in observe (str "type_of_lemma := " ++ Printer.pr_lconstr_env (Global.env ()) !evd type_of_lemma); type_of_lemma,type_info diff --git a/plugins/funind/recdef.ml b/plugins/funind/recdef.ml index 4fd9e0ff8..12ed758ba 100644 --- a/plugins/funind/recdef.ml +++ b/plugins/funind/recdef.ml @@ -657,7 +657,7 @@ let terminate_letin (na,b,t,e) expr_info continuation_tac info = continuation_tac {info with info = new_e; forbidden_ids = new_forbidden} let pf_type c tac gl = - let evars, ty = Typing.type_of (pf_env gl) (project gl) c in + let evars, ty = Typing.type_of (pf_env gl) (project gl) (EConstr.of_constr c) in tclTHEN (Refiner.tclEVARS evars) (tac ty) gl let pf_typel l tac = diff --git a/plugins/setoid_ring/newring.ml b/plugins/setoid_ring/newring.ml index e1b95ddbc..b0a3e839b 100644 --- a/plugins/setoid_ring/newring.ml +++ b/plugins/setoid_ring/newring.ml @@ -493,8 +493,8 @@ let ring_equality env evd (r,add,mul,opp,req) = match opp with Some opp -> plapp evd coq_eq_morph [|r;add;mul;opp|] | None -> plapp evd coq_eq_smorph [|r;add;mul|] in - let setoid = Typing.e_solve_evars env evd setoid in - let op_morph = Typing.e_solve_evars env evd op_morph in + let setoid = Typing.e_solve_evars env evd (EConstr.of_constr setoid) in + let op_morph = Typing.e_solve_evars env evd (EConstr.of_constr op_morph) in (setoid,op_morph) | _ -> let setoid = setoid_of_relation (Global.env ()) evd r req in @@ -581,7 +581,7 @@ let make_hyp_list env evd lH = (fun c l -> plapp evd coq_cons [|carrier; (make_hyp env evd c); l|]) lH (plapp evd coq_nil [|carrier|]) in - let l' = Typing.e_solve_evars env evd l in + let l' = Typing.e_solve_evars env evd (EConstr.of_constr l) in Evarutil.nf_evars_universes !evd l' let interp_power env evd pow = @@ -707,7 +707,7 @@ let make_term_list env evd carrier rl = let l = List.fold_right (fun x l -> plapp evd coq_cons [|carrier;x;l|]) rl (plapp evd coq_nil [|carrier|]) - in Typing.e_solve_evars env evd l + in Typing.e_solve_evars env evd (EConstr.of_constr l) let carg = Tacinterp.Value.of_constr let tacarg expr = diff --git a/pretyping/cases.ml b/pretyping/cases.ml index 04f50d50e..882c052f6 100644 --- a/pretyping/cases.ml +++ b/pretyping/cases.ml @@ -1381,7 +1381,7 @@ and match_current pb (initial,tomatch) = let case = make_case_or_project pb.env indf ci pred current brvals in - Typing.check_allowed_sort pb.env !(pb.evdref) mind current pred; + Typing.check_allowed_sort pb.env !(pb.evdref) mind (EConstr.of_constr current) (EConstr.of_constr pred); { uj_val = applist (case, inst); uj_type = prod_applist typ inst } @@ -1684,7 +1684,7 @@ let build_tycon loc env tycon_env s subst tycon extenv evdref t = (lift (n'-n) impossible_case_type, mkSort u) | Some t -> let t = abstract_tycon loc tycon_env evdref subst tycon extenv t in - let evd,tt = Typing.type_of extenv !evdref t in + let evd,tt = Typing.type_of extenv !evdref (EConstr.of_constr t) in evdref := evd; (t,tt) in let b = e_cumul env evdref (EConstr.of_constr tt) (EConstr.mkSort s) (* side effect *) in @@ -1920,7 +1920,7 @@ let prepare_predicate_from_arsign_tycon env sigma loc tomatchs arsign c = assert (len == 0); let p = predicate 0 c in let env' = List.fold_right push_rel_context arsign env in - try let sigma' = fst (Typing.type_of env' sigma p) in + try let sigma' = fst (Typing.type_of env' sigma (EConstr.of_constr p)) in Some (sigma', p) with e when precatchable_exception e -> None @@ -2041,7 +2041,7 @@ let constr_of_pat env evdref arsign pat avoid = let IndType (indf, _) = try find_rectype env ( !evdref) (EConstr.of_constr (lift (-(List.length realargs)) ty)) with Not_found -> error_case_not_inductive env !evdref - {uj_val = ty; uj_type = Typing.unsafe_type_of env !evdref ty} + {uj_val = ty; uj_type = Typing.unsafe_type_of env !evdref (EConstr.of_constr ty)} in let (ind,u), params = dest_ind_family indf in if not (eq_ind ind cind) then error_bad_constructor ~loc:l env cstr ind; @@ -2242,7 +2242,7 @@ let constrs_of_pats typing_fun env evdref eqns tomatchs sign neqs arity = let j = typing_fun (mk_tycon (EConstr.of_constr tycon)) rhs_env eqn.rhs.it in let bbody = it_mkLambda_or_LetIn j.uj_val rhs_rels' and btype = it_mkProd_or_LetIn j.uj_type rhs_rels' in - let _btype = evd_comb1 (Typing.type_of env) evdref bbody in + let _btype = evd_comb1 (Typing.type_of env) evdref (EConstr.of_constr bbody) in let branch_name = Id.of_string ("program_branch_" ^ (string_of_int !i)) in let branch_decl = LocalDef (Name branch_name, lift !i bbody, lift !i btype) in let branch = diff --git a/pretyping/coercion.ml b/pretyping/coercion.ml index 0ea6758a7..04e235cc5 100644 --- a/pretyping/coercion.ml +++ b/pretyping/coercion.ml @@ -188,7 +188,7 @@ and coerce loc env evdref (x : Term.constr) (y : Term.constr) aux (hdy :: tele) (subst1 hdx restT) (subst1 hdy restT') (succ i) (fun x -> eq_app (co x)) else Some (fun x -> - let term = co x in + let term = EConstr.of_constr (co x) in Typing.e_solve_evars env evdref term) in if isEvar c || isEvar c' || not (Program.is_program_generalized_coercion ()) then @@ -297,16 +297,16 @@ and coerce loc env evdref (x : Term.constr) (y : Term.constr) let evm = !evdref in (try subco () with NoSubtacCoercion -> - let typ = Typing.unsafe_type_of env evm c in - let typ' = Typing.unsafe_type_of env evm c' in + let typ = Typing.unsafe_type_of env evm (EConstr.of_constr c) in + let typ' = Typing.unsafe_type_of env evm (EConstr.of_constr c') in coerce_application typ typ' c c' l l') else subco () | x, y when Constr.equal c c' -> if Int.equal (Array.length l) (Array.length l') then let evm = !evdref in - let lam_type = Typing.unsafe_type_of env evm c in - let lam_type' = Typing.unsafe_type_of env evm c' in + let lam_type = Typing.unsafe_type_of env evm (EConstr.of_constr c) in + let lam_type' = Typing.unsafe_type_of env evm (EConstr.of_constr c') in coerce_application lam_type lam_type' c c' l l' else subco () | _ -> subco ()) @@ -337,7 +337,7 @@ let app_coercion env evdref coercion v = match coercion with | None -> v | Some f -> - let v' = Typing.e_solve_evars env evdref (f v) in + let v' = Typing.e_solve_evars env evdref (EConstr.of_constr (f v)) in whd_betaiota !evdref (EConstr.of_constr v') let coerce_itf loc env evd v t c1 = diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml index 570f95324..28ba60812 100644 --- a/pretyping/pretyping.ml +++ b/pretyping/pretyping.ml @@ -507,7 +507,7 @@ let pretype_ref loc evdref env ref us = | ref -> let evd, c = pretype_global loc univ_flexible env !evdref ref us in let () = evdref := evd in - let ty = Typing.unsafe_type_of env.ExtraEnv.env evd c in + let ty = Typing.unsafe_type_of env.ExtraEnv.env evd (EConstr.of_constr c) in make_judge c ty let judge_of_Type loc evd s = @@ -644,7 +644,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre { uj_val = it_mkLambda_or_LetIn j.uj_val ctxt; uj_type = it_mkProd_or_LetIn j.uj_type ctxt }) ctxtv vdef in - Typing.check_type_fixpoint loc env.ExtraEnv.env evdref names ftys vdefj; + Typing.check_type_fixpoint loc env.ExtraEnv.env evdref names (Array.map EConstr.of_constr ftys) vdefj; let ftys = Array.map (nf_evar !evdref) ftys in let fdefs = Array.map (fun x -> nf_evar !evdref (j_val x)) vdefj in let fixj = match fixkind with @@ -898,7 +898,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre let fj = pretype (mk_tycon (EConstr.of_constr fty)) env_f evdref lvar d in let v = let ind,_ = dest_ind_family indf in - Typing.check_allowed_sort env.ExtraEnv.env !evdref ind cj.uj_val p; + Typing.check_allowed_sort env.ExtraEnv.env !evdref ind (EConstr.of_constr cj.uj_val) (EConstr.of_constr p); obj ind p cj.uj_val fj.uj_val in { uj_val = v; uj_type = substl (realargs@[cj.uj_val]) ccl } @@ -917,7 +917,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre let p = it_mkLambda_or_LetIn (lift (nar+1) ccl) psign in let v = let ind,_ = dest_ind_family indf in - Typing.check_allowed_sort env.ExtraEnv.env !evdref ind cj.uj_val p; + Typing.check_allowed_sort env.ExtraEnv.env !evdref ind (EConstr.of_constr cj.uj_val) (EConstr.of_constr p); obj ind p cj.uj_val fj.uj_val in { uj_val = v; uj_type = ccl }) @@ -981,7 +981,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre let ind,_ = dest_ind_family indf in let ci = make_case_info env.ExtraEnv.env (fst ind) IfStyle in let pred = nf_evar !evdref pred in - Typing.check_allowed_sort env.ExtraEnv.env !evdref ind cj.uj_val pred; + Typing.check_allowed_sort env.ExtraEnv.env !evdref ind (EConstr.of_constr cj.uj_val) (EConstr.of_constr pred); mkCase (ci, pred, cj.uj_val, [|b1;b2|]) in let cj = { uj_val = v; uj_type = p } in diff --git a/pretyping/tacred.ml b/pretyping/tacred.ml index 290d77b1b..a3983737d 100644 --- a/pretyping/tacred.ml +++ b/pretyping/tacred.ml @@ -1164,7 +1164,7 @@ let pattern_occs loccs_trm = { e_redfun = begin fun env sigma c -> let sigma = Sigma.to_evar_map sigma in let abstr_trm, sigma = List.fold_right (abstract_scheme env) loccs_trm (EConstr.Unsafe.to_constr c,sigma) in try - let _ = Typing.unsafe_type_of env sigma abstr_trm in + let _ = Typing.unsafe_type_of env sigma (EConstr.of_constr abstr_trm) in Sigma.Unsafe.of_pair (applist(abstr_trm, List.map snd loccs_trm), sigma) with Type_errors.TypeError (env',t) -> raise (ReductionTacticError (InvalidAbstraction (env,sigma,abstr_trm,(env',t)))) diff --git a/pretyping/typing.ml b/pretyping/typing.ml index 64264cf08..c948f9b9a 100644 --- a/pretyping/typing.ml +++ b/pretyping/typing.ml @@ -20,6 +20,11 @@ open Typeops open Arguments_renaming open Context.Rel.Declaration +let push_rec_types pfix env = + let (i, c, t) = pfix in + let inj c = EConstr.Unsafe.to_constr c in + push_rec_types (i, Array.map inj c, Array.map inj t) env + let meta_type evd mv = let ty = try Evd.meta_ftype evd mv @@ -28,12 +33,12 @@ let meta_type evd mv = let constant_type_knowing_parameters env cst jl = let paramstyp = Array.map (fun j -> lazy j.uj_type) jl in - type_of_constant_knowing_parameters_in env cst paramstyp + EConstr.of_constr (type_of_constant_knowing_parameters_in env cst paramstyp) let inductive_type_knowing_parameters env (ind,u) jl = let mspec = lookup_mind_specif env ind in let paramstyp = Array.map (fun j -> lazy j.uj_type) jl in - Inductive.type_of_inductive_knowing_parameters env (mspec,u) paramstyp + EConstr.of_constr (Inductive.type_of_inductive_knowing_parameters env (mspec,u) paramstyp) let e_type_judgment env evdref j = match EConstr.kind !evdref (EConstr.of_constr (whd_all env !evdref (EConstr.of_constr j.uj_type))) with @@ -44,7 +49,7 @@ let e_type_judgment env evdref j = | _ -> error_not_type env j let e_assumption_of_judgment env evdref j = - try (e_type_judgment env evdref j).utj_val + try EConstr.of_constr (e_type_judgment env evdref j).utj_val with TypeError _ -> error_assumption env j @@ -84,27 +89,28 @@ let max_sort l = if Sorts.List.mem InSet l then InSet else InProp let e_is_correct_arity env evdref c pj ind specif params = + let open EConstr in let arsign = make_arity_signature env true (make_ind_family (ind,params)) in let allowed_sorts = elim_sorts specif in let error () = error_elim_arity env ind allowed_sorts c pj None in let rec srec env pt ar = - let pt' = whd_all env !evdref (EConstr.of_constr pt) in - match kind_of_term pt', ar with + let pt' = EConstr.of_constr (whd_all env !evdref pt) in + match EConstr.kind !evdref pt', ar with | Prod (na1,a1,t), (LocalAssum (_,a1'))::ar' -> - if not (Evarconv.e_cumul env evdref (EConstr.of_constr a1) (EConstr.of_constr a1')) then error (); - srec (push_rel (LocalAssum (na1,a1)) env) t ar' + if not (Evarconv.e_cumul env evdref a1 (EConstr.of_constr a1')) then error (); + srec (push_rel (LocalAssum (na1,EConstr.Unsafe.to_constr a1)) env) t ar' | Sort s, [] -> if not (Sorts.List.mem (Sorts.family s) allowed_sorts) then error () | Evar (ev,_), [] -> let evd, s = Evd.fresh_sort_in_family env !evdref (max_sort allowed_sorts) in - evdref := Evd.define ev (mkSort s) evd + evdref := Evd.define ev (Constr.mkSort s) evd | _, (LocalDef _ as d)::ar' -> - srec (push_rel d env) (lift 1 pt') ar' + srec (push_rel d env) (Vars.lift 1 pt') ar' | _ -> error () in - srec env pj.uj_type (List.rev arsign) + srec env (EConstr.of_constr pj.uj_type) (List.rev arsign) let e_type_case_branches env evdref (ind,largs) pj c = let specif = lookup_mind_specif env (fst ind) in @@ -128,24 +134,25 @@ let e_judge_of_case env evdref ci pj cj lfj = uj_type = rslty } let check_type_fixpoint loc env evdref lna lar vdefj = + let open EConstr in let lt = Array.length vdefj in if Int.equal (Array.length lar) lt then for i = 0 to lt-1 do if not (Evarconv.e_cumul env evdref (EConstr.of_constr (vdefj.(i)).uj_type) - (EConstr.of_constr (lift lt lar.(i)))) then + (Vars.lift lt lar.(i))) then Pretype_errors.error_ill_typed_rec_body ~loc env !evdref - i lna vdefj lar + i lna vdefj (Array.map EConstr.Unsafe.to_constr lar) done (* FIXME: might depend on the level of actual parameters!*) let check_allowed_sort env sigma ind c p = - let pj = Retyping.get_judgment_of env sigma (EConstr.of_constr p) in + let pj = Retyping.get_judgment_of env sigma p in let ksort = family_of_sort (sort_of_arity env sigma (EConstr.of_constr pj.uj_type)) in let specif = Global.lookup_inductive (fst ind) in let sorts = elim_sorts specif in if not (List.exists ((==) ksort) sorts) then let s = inductive_sort_family (snd specif) in - error_elim_arity env ind sorts c pj + error_elim_arity env ind sorts (EConstr.Unsafe.to_constr c) pj (Some(ksort,s,error_elim_explain ksort s)) let e_judge_of_cast env evdref cj k tj = @@ -160,21 +167,36 @@ let enrich_env env evdref = let penv' = Pre_env.({ penv with env_stratification = { penv.env_stratification with env_universes = Evd.universes !evdref } }) in Environ.env_of_pre_env penv' + +let check_fix env sigma pfix = + let inj c = EConstr.to_constr sigma c in + let (idx, (ids, cs, ts)) = pfix in + check_fix env (idx, (ids, Array.map inj cs, Array.map inj ts)) + +let check_cofix env sigma pcofix = + let inj c = EConstr.to_constr sigma c in + let (idx, (ids, cs, ts)) = pcofix in + check_cofix env (idx, (ids, Array.map inj cs, Array.map inj ts)) + +let make_judge c ty = + make_judge (EConstr.Unsafe.to_constr c) (EConstr.Unsafe.to_constr ty) (* The typing machine with universes and existential variables. *) (* cstr must be in n.f. w.r.t. evars and execute returns a judgement where both the term and type are in n.f. *) let rec execute env evdref cstr = - match kind_of_term cstr with + let open EConstr in + let cstr = EConstr.of_constr (whd_evar !evdref (EConstr.Unsafe.to_constr cstr)) in + match EConstr.kind !evdref cstr with | Meta n -> - { uj_val = cstr; uj_type = meta_type !evdref n } + { uj_val = EConstr.Unsafe.to_constr cstr; uj_type = meta_type !evdref n } | Evar ev -> - let ty = Evd.existential_type !evdref ev in - let jty = execute env evdref (whd_evar !evdref ty) in + let ty = EConstr.existential_type !evdref ev in + let jty = execute env evdref ty in let jty = e_assumption_of_judgment env evdref jty in - { uj_val = cstr; uj_type = jty } + { uj_val = EConstr.Unsafe.to_constr cstr; uj_type = EConstr.Unsafe.to_constr jty } | Rel n -> judge_of_relative env n @@ -183,13 +205,13 @@ let rec execute env evdref cstr = judge_of_variable env id | Const c -> - make_judge cstr (rename_type_of_constant env c) + make_judge cstr (EConstr.of_constr (rename_type_of_constant env c)) | Ind ind -> - make_judge cstr (rename_type_of_inductive env ind) + make_judge cstr (EConstr.of_constr (rename_type_of_inductive env ind)) | Construct cstruct -> - make_judge cstr (rename_type_of_constructor env cstruct) + make_judge cstr (EConstr.of_constr (rename_type_of_constructor env cstruct)) | Case (ci,p,c,lf) -> let cj = execute env evdref c in @@ -200,13 +222,13 @@ let rec execute env evdref cstr = | Fix ((vn,i as vni),recdef) -> let (_,tys,_ as recdef') = execute_recdef env evdref recdef in let fix = (vni,recdef') in - check_fix env fix; + check_fix env !evdref fix; make_judge (mkFix fix) tys.(i) | CoFix (i,recdef) -> let (_,tys,_ as recdef') = execute_recdef env evdref recdef in let cofix = (i,recdef') in - check_cofix env cofix; + check_cofix env !evdref cofix; make_judge (mkCoFix cofix) tys.(i) | Sort (Prop c) -> @@ -222,7 +244,7 @@ let rec execute env evdref cstr = | App (f,args) -> let jl = execute_array env evdref args in let j = - match kind_of_term f with + match EConstr.kind !evdref f with | Ind ind when Environ.template_polymorphic_pind ind env -> (* Sort-polymorphism of inductive types *) make_judge f @@ -273,7 +295,7 @@ and execute_recdef env evdref (names,lar,vdef) = let lara = Array.map (e_assumption_of_judgment env evdref) larj in let env1 = push_rec_types (names,lara,vdef) env in let vdefj = execute_array env1 evdref vdef in - let vdefv = Array.map j_val vdefj in + let vdefv = Array.map (j_val %> EConstr.of_constr) vdefj in let _ = check_type_fixpoint Loc.ghost env1 evdref names lara vdefj in (names,lara,vdefv) @@ -282,8 +304,8 @@ and execute_array env evdref = Array.map (execute env evdref) let e_check env evdref c t = let env = enrich_env env evdref in let j = execute env evdref c in - if not (Evarconv.e_cumul env evdref (EConstr.of_constr j.uj_type) (EConstr.of_constr t)) then - error_actual_type env j (nf_evar !evdref t) + if not (Evarconv.e_cumul env evdref (EConstr.of_constr j.uj_type) t) then + error_actual_type env j (EConstr.to_constr !evdref t) (* Type of a constr *) @@ -328,4 +350,4 @@ let e_solve_evars env evdref c = (* side-effect on evdref *) nf_evar !evdref c -let _ = Evarconv.set_solve_evars (fun env evdref c -> EConstr.of_constr (e_solve_evars env evdref (EConstr.Unsafe.to_constr c))) +let _ = Evarconv.set_solve_evars (fun env evdref c -> EConstr.of_constr (e_solve_evars env evdref c)) diff --git a/pretyping/typing.mli b/pretyping/typing.mli index 04e5e40bc..3c1c4324d 100644 --- a/pretyping/typing.mli +++ b/pretyping/typing.mli @@ -14,33 +14,33 @@ open Evd and universes. *) (** Typecheck a term and return its type. May trigger an evarmap leak. *) -val unsafe_type_of : env -> evar_map -> constr -> types +val unsafe_type_of : env -> evar_map -> EConstr.constr -> types (** Typecheck a term and return its type + updated evars, optionally refreshing universes *) -val type_of : ?refresh:bool -> env -> evar_map -> constr -> evar_map * types +val type_of : ?refresh:bool -> env -> evar_map -> EConstr.constr -> evar_map * types (** Variant of [type_of] using references instead of state-passing. *) -val e_type_of : ?refresh:bool -> env -> evar_map ref -> constr -> types +val e_type_of : ?refresh:bool -> env -> evar_map ref -> EConstr.constr -> types (** Typecheck a type and return its sort *) -val e_sort_of : env -> evar_map ref -> types -> sorts +val e_sort_of : env -> evar_map ref -> EConstr.types -> sorts (** Typecheck a term has a given type (assuming the type is OK) *) -val e_check : env -> evar_map ref -> constr -> types -> unit +val e_check : env -> evar_map ref -> EConstr.constr -> EConstr.types -> unit (** Returns the instantiated type of a metavariable *) val meta_type : evar_map -> metavariable -> types (** Solve existential variables using typing *) -val e_solve_evars : env -> evar_map ref -> constr -> constr +val e_solve_evars : env -> evar_map ref -> EConstr.constr -> constr (** Raise an error message if incorrect elimination for this inductive *) (** (first constr is term to match, second is return predicate) *) -val check_allowed_sort : env -> evar_map -> pinductive -> constr -> constr -> +val check_allowed_sort : env -> evar_map -> pinductive -> EConstr.constr -> EConstr.constr -> unit (** Raise an error message if bodies have types not unifiable with the expected ones *) val check_type_fixpoint : Loc.t -> env -> evar_map ref -> - Names.Name.t array -> types array -> unsafe_judgment array -> unit + Names.Name.t array -> EConstr.types array -> unsafe_judgment array -> unit diff --git a/pretyping/unification.ml b/pretyping/unification.ml index ac2f14051..f418dc6a9 100644 --- a/pretyping/unification.ml +++ b/pretyping/unification.ml @@ -102,7 +102,7 @@ let abstract_list_all env evd typ c l = let l_with_all_occs = List.map (function a -> (LikeFirst,a)) l in let p,evd = abstract_scheme env evd c l_with_all_occs ctxt in let evd,typp = - try Typing.type_of env evd p + try Typing.type_of env evd (EConstr.of_constr p) with | UserError _ -> error_cannot_find_well_typed_abstraction env evd p (List.map EConstr.of_constr l) None @@ -1214,7 +1214,7 @@ let applyHead env (type r) (evd : r Sigma.t) n c = apprec (n-1) (mkApp(c,[|evar|])) (subst1 evar c2) (p +> q) evd' | _ -> error "Apply_Head_Then" in - apprec n c (Typing.unsafe_type_of env (Sigma.to_evar_map evd) c) Sigma.refl evd + apprec n c (Typing.unsafe_type_of env (Sigma.to_evar_map evd) (EConstr.of_constr c)) Sigma.refl evd let is_mimick_head ts f = match kind_of_term f with diff --git a/proofs/clenv.ml b/proofs/clenv.ml index f4ac094b8..c2130a64a 100644 --- a/proofs/clenv.ml +++ b/proofs/clenv.ml @@ -29,7 +29,7 @@ open Sigma.Notations (* Abbreviations *) let pf_env = Refiner.pf_env -let pf_type_of gls c = Typing.unsafe_type_of (pf_env gls) gls.sigma c +let pf_type_of gls c = Typing.unsafe_type_of (pf_env gls) gls.sigma (EConstr.of_constr c) (******************************************************************) (* Clausal environments *) diff --git a/proofs/logic.ml b/proofs/logic.ml index 2df626e1c..93b31ced1 100644 --- a/proofs/logic.ml +++ b/proofs/logic.ml @@ -86,7 +86,7 @@ let apply_to_hyp check sign id f = else sign let check_typability env sigma c = - if !check then let _ = unsafe_type_of env sigma c in () + if !check then let _ = unsafe_type_of env sigma (EConstr.of_constr c) in () (************************************************************************) (************************************************************************) @@ -330,7 +330,7 @@ let meta_free_prefix sigma a = with Stop acc -> Array.rev_of_list acc let goal_type_of env sigma c = - if !check then unsafe_type_of env sigma c + if !check then unsafe_type_of env sigma (EConstr.of_constr c) else Retyping.get_type_of env sigma (EConstr.of_constr c) let rec mk_refgoals sigma goal goalacc conclty trm = diff --git a/proofs/refine.ml b/proofs/refine.ml index e6e3ef47d..b62f0bea4 100644 --- a/proofs/refine.ml +++ b/proofs/refine.ml @@ -28,12 +28,12 @@ let typecheck_evar ev env sigma = let info = Evd.find sigma ev in (** Typecheck the hypotheses. *) let type_hyp (sigma, env) decl = - let t = NamedDecl.get_type decl in + let t = EConstr.of_constr (NamedDecl.get_type decl) in let evdref = ref sigma in let _ = Typing.e_sort_of env evdref t in let () = match decl with | LocalAssum _ -> () - | LocalDef (_,body,_) -> Typing.e_check env evdref body t + | LocalDef (_,body,_) -> Typing.e_check env evdref (EConstr.of_constr body) t in (!evdref, Environ.push_named decl env) in @@ -42,12 +42,12 @@ let typecheck_evar ev env sigma = let (sigma, env) = List.fold_left type_hyp (sigma, env) changed in (** Typecheck the conclusion *) let evdref = ref sigma in - let _ = Typing.e_sort_of env evdref (Evd.evar_concl info) in + let _ = Typing.e_sort_of env evdref (EConstr.of_constr (Evd.evar_concl info)) in !evdref let typecheck_proof c concl env sigma = let evdref = ref sigma in - let () = Typing.e_check env evdref c concl in + let () = Typing.e_check env evdref (EConstr.of_constr c) (EConstr.of_constr concl) in !evdref let (pr_constrv,pr_constr) = diff --git a/proofs/tacmach.ml b/proofs/tacmach.ml index b63b2ce28..148f9d675 100644 --- a/proofs/tacmach.ml +++ b/proofs/tacmach.ml @@ -93,8 +93,8 @@ let pf_nf = pf_reduce' simpl let pf_nf_betaiota = pf_reduce' (fun _ -> nf_betaiota) let pf_compute = pf_reduce' compute let pf_unfoldn ubinds = pf_reduce' (unfoldn ubinds) -let pf_unsafe_type_of = pf_reduce unsafe_type_of -let pf_type_of = pf_reduce type_of +let pf_unsafe_type_of = pf_reduce' unsafe_type_of +let pf_type_of = pf_reduce' type_of let pf_get_type_of = pf_reduce Retyping.get_type_of let pf_conv_x gl = pf_apply test_conversion gl Reduction.CONV @@ -175,10 +175,10 @@ module New = struct let pf_concl = Proofview.Goal.concl let pf_unsafe_type_of gl t = - pf_apply unsafe_type_of gl t + pf_apply unsafe_type_of gl (EConstr.of_constr t) let pf_type_of gl t = - pf_apply type_of gl t + pf_apply type_of gl (EConstr.of_constr t) let pf_conv_x gl t1 t2 = pf_apply is_conv gl t1 t2 diff --git a/tactics/autorewrite.ml b/tactics/autorewrite.ml index 46600cdd7..80b9ec06e 100644 --- a/tactics/autorewrite.ml +++ b/tactics/autorewrite.ml @@ -272,7 +272,7 @@ let decompose_applied_relation metas env sigma c ctype left2right = try let others,(c1,c2) = split_last_two args in let ty1, ty2 = - Typing.unsafe_type_of env eqclause.evd c1, Typing.unsafe_type_of env eqclause.evd c2 + Typing.unsafe_type_of env eqclause.evd (EConstr.of_constr c1), Typing.unsafe_type_of env eqclause.evd (EConstr.of_constr c2) in (* if not (evd_convertible env eqclause.evd ty1 ty2) then None *) (* else *) @@ -290,7 +290,7 @@ let decompose_applied_relation metas env sigma c ctype left2right = | None -> None let find_applied_relation metas loc env sigma c left2right = - let ctype = Typing.unsafe_type_of env sigma c in + let ctype = Typing.unsafe_type_of env sigma (EConstr.of_constr c) in match decompose_applied_relation metas env sigma c ctype left2right with | Some c -> c | None -> diff --git a/tactics/equality.ml b/tactics/equality.ml index 17038e42d..58c86ff42 100644 --- a/tactics/equality.ml +++ b/tactics/equality.ml @@ -1085,7 +1085,7 @@ let find_sigma_data env s = build_sigma_type () let make_tuple env sigma (rterm,rty) lind = assert (not (EConstr.Vars.noccurn sigma lind (EConstr.of_constr rty))); let sigdata = find_sigma_data env (get_sort_of env sigma (EConstr.of_constr rty)) in - let sigma, a = type_of ~refresh:true env sigma (mkRel lind) in + let sigma, a = type_of ~refresh:true env sigma (EConstr.mkRel lind) in let na = Context.Rel.Declaration.get_name (lookup_rel lind env) in (* We move [lind] to [1] and lift other rels > [lind] by 1 *) let rty = lift (1-lind) (liftn lind (lind+1) rty) in @@ -1119,7 +1119,7 @@ let minimal_free_rels_rec env sigma = let rec minimalrec_free_rels_rec prev_rels (c,cty) = let (cty,direct_rels) = minimal_free_rels env sigma (c,cty) in let combined_rels = Int.Set.union prev_rels direct_rels in - let folder rels i = snd (minimalrec_free_rels_rec rels (c, unsafe_type_of env sigma (mkRel i))) + let folder rels i = snd (minimalrec_free_rels_rec rels (c, unsafe_type_of env sigma (EConstr.mkRel i))) in (cty, List.fold_left folder combined_rels (Int.Set.elements (Int.Set.diff direct_rels prev_rels))) in minimalrec_free_rels_rec Int.Set.empty @@ -1165,7 +1165,7 @@ let sig_clausal_form env sigma sort_of_ty siglen ty dflt = let rec sigrec_clausal_form siglen p_i = if Int.equal siglen 0 then (* is the default value typable with the expected type *) - let dflt_typ = unsafe_type_of env sigma dflt in + let dflt_typ = unsafe_type_of env sigma (EConstr.of_constr dflt) in try let () = evdref := Evarconv.the_conv_x_leq env (EConstr.of_constr dflt_typ) (EConstr.of_constr p_i) !evdref in let () = evdref := Evarconv.consider_remaining_unif_problems env !evdref in @@ -1184,7 +1184,7 @@ let sig_clausal_form env sigma sort_of_ty siglen ty dflt = (destEvar ev) with | Some w -> - let w_type = unsafe_type_of env sigma w in + let w_type = unsafe_type_of env sigma (EConstr.of_constr w) in if Evarconv.e_cumul env evdref (EConstr.of_constr w_type) (EConstr.of_constr a) then let exist_term = Evarutil.evd_comb1 (Evd.fresh_global env) evdref sigdata.intro in applist(exist_term,[a;p_i_minus_1;w;tuple_tail]) @@ -1273,7 +1273,7 @@ let make_iterated_tuple env sigma dflt (z,zty) = sigma, (tuple,tuplety,dfltval) let rec build_injrec env sigma dflt c = function - | [] -> make_iterated_tuple env sigma dflt (c,unsafe_type_of env sigma c) + | [] -> make_iterated_tuple env sigma dflt (c,unsafe_type_of env sigma (EConstr.of_constr c)) | ((sp,cnum),argnum)::l -> try let (cnum_nlams,cnum_env,kont) = descend_then env sigma c cnum in @@ -1367,7 +1367,7 @@ let inject_at_positions env sigma l2r (eq,_,(t,t1,t2)) eq_clause posns tac = let injfun = mkNamedLambda e t injbody in let sigma,congr = Evd.fresh_global env sigma eq.congr in let pf = applist(congr,[t;resty;injfun;t1;t2]) in - let sigma, pf_typ = Typing.type_of env sigma pf in + let sigma, pf_typ = Typing.type_of env sigma (EConstr.of_constr pf) in let inj_clause = apply_on_clause (pf,pf_typ) eq_clause in let pf = Clenvtac.clenv_value_cast_meta inj_clause in let ty = simplify_args env sigma (clenv_type inj_clause) in @@ -1555,8 +1555,8 @@ let subst_tuple_term env sigma dep_pair1 dep_pair2 b = (* Simulate now the normalisation treatment made by Logic.mk_refgoals *) let expected_goal = nf_betaiota sigma (EConstr.of_constr expected_goal) in (* Retype to get universes right *) - let sigma, expected_goal_ty = Typing.type_of env sigma expected_goal in - let sigma, _ = Typing.type_of env sigma body in + let sigma, expected_goal_ty = Typing.type_of env sigma (EConstr.of_constr expected_goal) in + let sigma, _ = Typing.type_of env sigma (EConstr.of_constr body) in Sigma.Unsafe.of_pair ((body, expected_goal), sigma) (* Like "replace" but decompose dependent equalities *) diff --git a/tactics/hints.ml b/tactics/hints.ml index 2aa434777..63d10573a 100644 --- a/tactics/hints.ml +++ b/tactics/hints.ml @@ -910,7 +910,7 @@ let make_mode ref m = let make_trivial env sigma poly ?(name=PathAny) r = let c,ctx = fresh_global_or_constr env sigma poly r in let sigma = Evd.merge_context_set univ_flexible sigma ctx in - let t = hnf_constr env sigma (EConstr.of_constr (unsafe_type_of env sigma c)) in + let t = hnf_constr env sigma (EConstr.of_constr (unsafe_type_of env sigma (EConstr.of_constr c))) in let hd = head_of_constr_reference (head_constr sigma t) in let ce = mk_clenv_from_env env sigma None (c,t) in (Some hd, { pri=1; diff --git a/tactics/inv.ml b/tactics/inv.ml index 38f75995b..9282af759 100644 --- a/tactics/inv.ml +++ b/tactics/inv.ml @@ -123,13 +123,13 @@ let make_inv_predicate env evd indf realargs id status concl = let refl_term = eqdata.Coqlib.refl in let refl_term = Evarutil.evd_comb1 (Evd.fresh_global env) evd refl_term in let refl = mkApp (refl_term, [|eqnty; rhs|]) in - let _ = Evarutil.evd_comb1 (Typing.type_of env) evd refl in + let _ = Evarutil.evd_comb1 (Typing.type_of env) evd (EConstr.of_constr refl) in let args = refl :: args in build_concl eqns args (succ n) restlist in let (newconcl, args) = build_concl [] [] 0 realargs in let predicate = it_mkLambda_or_LetIn_name env newconcl hyps in - let _ = Evarutil.evd_comb1 (Typing.type_of env) evd predicate in + let _ = Evarutil.evd_comb1 (Typing.type_of env) evd (EConstr.of_constr predicate) in (* OK - this predicate should now be usable by res_elimination_then to do elimination on the conclusion. *) predicate, args diff --git a/tactics/tactics.ml b/tactics/tactics.ml index c2163a274..8fb47b994 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -210,7 +210,7 @@ let convert_concl ?(check=true) ty k = let Sigma ((), sigma, p) = if check then begin let sigma = Sigma.to_evar_map sigma in - ignore (Typing.unsafe_type_of env sigma ty); + ignore (Typing.unsafe_type_of env sigma (EConstr.of_constr ty)); let sigma,b = Reductionops.infer_conv env sigma ty conclty in if not b then error "Not convertible."; Sigma.Unsafe.of_pair ((), sigma) @@ -827,7 +827,7 @@ let change_on_subterm cv_pb deep t where = { e_redfun = begin fun env sigma c -> env sigma c in if !mayneedglobalcheck then begin - try ignore (Typing.unsafe_type_of env (Sigma.to_evar_map sigma) c) + try ignore (Typing.unsafe_type_of env (Sigma.to_evar_map sigma) (EConstr.of_constr c)) with e when catchable_exception e -> error "Replacement would lead to an ill-typed term." end; @@ -1228,7 +1228,7 @@ let cut c = let is_sort = try (** Backward compat: ensure that [c] is well-typed. *) - let typ = Typing.unsafe_type_of env sigma c in + let typ = Typing.unsafe_type_of env sigma (EConstr.of_constr c) in let typ = whd_all env sigma (EConstr.of_constr typ) in match kind_of_term typ with | Sort _ -> true @@ -1940,7 +1940,7 @@ let exact_check c = let concl = Proofview.Goal.concl (Proofview.Goal.assume gl) in let env = Proofview.Goal.env gl in let sigma = Sigma.to_evar_map sigma in - let sigma, ct = Typing.type_of env sigma c in + let sigma, ct = Typing.type_of env sigma (EConstr.of_constr c) in let tac = Tacticals.New.tclTHEN (convert_leq ct concl) (exact_no_check c) in @@ -2009,20 +2009,20 @@ exception DependsOnBody of Id.t option let check_is_type env sigma ty = let evdref = ref sigma in try - let _ = Typing.e_sort_of env evdref ty in + let _ = Typing.e_sort_of env evdref (EConstr.of_constr ty) in !evdref with e when CErrors.noncritical e -> raise (DependsOnBody None) let check_decl env sigma decl = let open Context.Named.Declaration in - let ty = NamedDecl.get_type decl in + let ty = EConstr.of_constr (NamedDecl.get_type decl) in let evdref = ref sigma in try let _ = Typing.e_sort_of env evdref ty in let _ = match decl with | LocalAssum _ -> () - | LocalDef (_,c,_) -> Typing.e_check env evdref c ty + | LocalDef (_,c,_) -> Typing.e_check env evdref (EConstr.of_constr c) ty in !evdref with e when CErrors.noncritical e -> @@ -2622,7 +2622,7 @@ let letin_tac_gen with_eq (id,depdecls,lastlhyp,ccl,c) ty = let refl = applist (refl, [t;mkVar id]) in let term = mkNamedLetIn id c t (mkLetIn (Name heq, refl, eq, ccl)) in let sigma = Sigma.to_evar_map sigma in - let sigma, _ = Typing.type_of env sigma term in + let sigma, _ = Typing.type_of env sigma (EConstr.of_constr term) in let ans = term, Tacticals.New.tclTHEN (intro_gen (NamingMustBe (loc,heq)) (decode_hyp lastlhyp) true false) @@ -2783,7 +2783,7 @@ let generalize_goal_gen env sigma ids i ((occs,c,b),na) t cl = let generalize_goal gl i ((occs,c,b),na as o) (cl,sigma) = let env = Tacmach.pf_env gl in let ids = Tacmach.pf_ids_of_hyps gl in - let sigma, t = Typing.type_of env sigma c in + let sigma, t = Typing.type_of env sigma (EConstr.of_constr c) in generalize_goal_gen env sigma ids i o t cl let old_generalize_dep ?(with_let=false) c gl = @@ -2818,7 +2818,7 @@ let old_generalize_dep ?(with_let=false) c gl = let cl'',evd = generalize_goal gl 0 ((AllOccurrences,c,body),Anonymous) (cl',project gl) in (** Check that the generalization is indeed well-typed *) - let (evd, _) = Typing.type_of env evd cl'' in + let (evd, _) = Typing.type_of env evd (EConstr.of_constr cl'') in let args = Context.Named.to_instance to_quantify_rev in tclTHENLIST [tclEVARS evd; @@ -2836,7 +2836,7 @@ let generalize_gen_let lconstr = Proofview.Goal.nf_s_enter { s_enter = begin fun List.fold_right_i (Tacmach.New.of_old generalize_goal gl) 0 lconstr (Tacmach.New.pf_concl gl,Tacmach.New.project gl) in - let (evd, _) = Typing.type_of env evd newcl in + let (evd, _) = Typing.type_of env evd (EConstr.of_constr newcl) in let map ((_, c, b),_) = if Option.is_empty b then Some c else None in let tac = apply_type newcl (List.map_filter map lconstr) in Sigma.Unsafe.of_pair (tac, evd) @@ -2853,7 +2853,7 @@ let new_generalize_gen_let lconstr = let newcl, sigma, args = List.fold_right_i (fun i ((_,c,b),_ as o) (cl, sigma, args) -> - let sigma, t = Typing.type_of env sigma c in + let sigma, t = Typing.type_of env sigma (EConstr.of_constr c) in let args = if Option.is_empty b then c :: args else args in let cl, sigma = generalize_goal_gen env sigma ids i o t cl in (cl, sigma, args)) @@ -4738,7 +4738,7 @@ let prove_transitivity hdcncl eq_kind t = let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in let type_of = Typing.unsafe_type_of env sigma in - let typt = type_of t in + let typt = type_of (EConstr.of_constr t) in (mkApp(hdcncl, [| typ1; c1; typt ;t |]), mkApp(hdcncl, [| typt; t; typ2; c2 |])) in diff --git a/toplevel/class.ml b/toplevel/class.ml index 5aa000b16..86fd4d9a2 100644 --- a/toplevel/class.ml +++ b/toplevel/class.ml @@ -206,7 +206,7 @@ let build_id_coercion idf_opt source poly = let _ = if not (Reductionops.is_conv_leq env sigma - (EConstr.of_constr (Typing.unsafe_type_of env sigma val_f)) (EConstr.of_constr typ_f)) + (EConstr.of_constr (Typing.unsafe_type_of env sigma (EConstr.of_constr val_f))) (EConstr.of_constr typ_f)) then user_err (strbrk "Cannot be defined as coercion (maybe a bad number of arguments).") diff --git a/toplevel/command.ml b/toplevel/command.ml index 14a45f837..c6dd3eb99 100644 --- a/toplevel/command.ml +++ b/toplevel/command.ml @@ -953,7 +953,7 @@ let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation = let binders = letbinders @ [arg] in let binders_env = push_rel_context binders_rel env in let rel, _ = interp_constr_evars_impls env evdref r in - let relty = Typing.unsafe_type_of env !evdref rel in + let relty = Typing.unsafe_type_of env !evdref (EConstr.of_constr rel) in let relargty = let error () = user_err ~loc:(constr_loc r) @@ -1034,7 +1034,7 @@ let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation = ~src:(Loc.ghost, Evar_kinds.QuestionMark (Evar_kinds.Define false)) wf_proof; prop |]) in - let def = Typing.e_solve_evars env evdref def in + let def = Typing.e_solve_evars env evdref (EConstr.of_constr def) in let _ = evdref := Evarutil.nf_evar_map !evdref in let def = mkApp (def, [|intern_body_lam|]) in let binders_rel = nf_evar_context !evdref binders_rel in @@ -1105,11 +1105,11 @@ let interp_recursive isfix fixl notations = List.fold_left2 (fun env' id t -> if Flags.is_program_mode () then - let sort = Evarutil.evd_comb1 (Typing.type_of ~refresh:true env) evdref t in + let sort = Evarutil.evd_comb1 (Typing.type_of ~refresh:true env) evdref (EConstr.of_constr t) in let fixprot = try let app = mkApp (delayed_force fix_proto, [|sort; t|]) in - Typing.e_solve_evars env evdref app + Typing.e_solve_evars env evdref (EConstr.of_constr app) with e when CErrors.noncritical e -> t in LocalAssum (id,fixprot) :: env' -- cgit v1.2.3 From 77e638121b6683047be915da9d0499a58fcb6e52 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sun, 6 Nov 2016 19:30:24 +0100 Subject: Patternops API using EConstr. --- ltac/tacinterp.ml | 4 ++-- plugins/quote/quote.ml | 4 ++-- pretyping/patternops.ml | 43 +++++++++++++++++++++++++++++-------------- pretyping/patternops.mli | 3 ++- tactics/hints.ml | 10 +++++----- 5 files changed, 40 insertions(+), 24 deletions(-) diff --git a/ltac/tacinterp.ml b/ltac/tacinterp.ml index 8f5ac7e76..85d4ac06e 100644 --- a/ltac/tacinterp.ml +++ b/ltac/tacinterp.ml @@ -691,7 +691,7 @@ let interp_pure_open_constr ist = let interp_typed_pattern ist env sigma (_,c,_) = let sigma, c = interp_gen WithoutTypeConstraint ist true pure_open_constr_flags env sigma c in - pattern_of_constr env sigma c + pattern_of_constr env sigma (EConstr.of_constr c) (* Interprets a constr expression *) let interp_constr_in_compound_list inj_fun dest_fun interp_fun ist env sigma l = @@ -736,7 +736,7 @@ let interp_closed_typed_pattern_with_occurrences ist env sigma (occs, a) = try Inl (coerce_to_evaluable_ref env x) with CannotCoerceTo _ -> let c = coerce_to_closed_constr env x in - Inr (pattern_of_constr env sigma c) in + Inr (pattern_of_constr env sigma (EConstr.of_constr c)) in (try try_interp_ltac_var coerce_eval_ref_or_constr ist (Some (env,sigma)) (loc,id) with Not_found -> error_global_not_found ~loc (qualid_of_ident id)) diff --git a/plugins/quote/quote.ml b/plugins/quote/quote.ml index afc7e6665..a13948f77 100644 --- a/plugins/quote/quote.ml +++ b/plugins/quote/quote.ml @@ -212,9 +212,9 @@ let compute_rhs bodyi index_of_f = let i = destRel (Array.last args) in PMeta (Some (coerce_meta_in i)) | App (f,args) -> - PApp (pattern_of_constr (Global.env()) Evd.empty f, Array.map aux args) + PApp (pattern_of_constr (Global.env()) Evd.empty (EConstr.of_constr f), Array.map aux args) | Cast (c,_,_) -> aux c - | _ -> pattern_of_constr (Global.env())(*FIXME*) Evd.empty c + | _ -> pattern_of_constr (Global.env())(*FIXME*) Evd.empty (EConstr.of_constr c) in aux bodyi diff --git a/pretyping/patternops.ml b/pretyping/patternops.ml index 938b6b18e..d473f41bd 100644 --- a/pretyping/patternops.ml +++ b/pretyping/patternops.ml @@ -114,17 +114,27 @@ let rec head_pattern_bound t = | PLambda _ -> raise BoundPattern | PCoFix _ -> anomaly ~label:"head_pattern_bound" (Pp.str "not a type") -let head_of_constr_reference c = match kind_of_term c with +let head_of_constr_reference sigma c = match EConstr.kind sigma c with | Const (sp,_) -> ConstRef sp | Construct (sp,_) -> ConstructRef sp | Ind (sp,_) -> IndRef sp | Var id -> VarRef id | _ -> anomaly (Pp.str "Not a rigid reference") +let local_assum (na, t) = + let open Context.Rel.Declaration in + let inj = EConstr.Unsafe.to_constr in + LocalAssum (na, inj t) + +let local_def (na, b, t) = + let open Context.Rel.Declaration in + let inj = EConstr.Unsafe.to_constr in + LocalDef (na, inj b, inj t) + let pattern_of_constr env sigma t = + let open EConstr in let rec pattern_of_constr env t = - let open Context.Rel.Declaration in - match kind_of_term t with + match EConstr.kind sigma t with | Rel n -> PRel n | Meta n -> PMeta (Some (Id.of_string ("META" ^ string_of_int n))) | Var id -> PVar id @@ -133,14 +143,14 @@ let pattern_of_constr env sigma t = | Sort (Type _) -> PSort (GType []) | Cast (c,_,_) -> pattern_of_constr env c | LetIn (na,c,t,b) -> PLetIn (na,pattern_of_constr env c, - pattern_of_constr (push_rel (LocalDef (na,c,t)) env) b) + pattern_of_constr (push_rel (local_def (na,c,t)) env) b) | Prod (na,c,b) -> PProd (na,pattern_of_constr env c, - pattern_of_constr (push_rel (LocalAssum (na, c)) env) b) + pattern_of_constr (push_rel (local_assum (na, c)) env) b) | Lambda (na,c,b) -> PLambda (na,pattern_of_constr env c, - pattern_of_constr (push_rel (LocalAssum (na, c)) env) b) + pattern_of_constr (push_rel (local_assum (na, c)) env) b) | App (f,a) -> (match - match kind_of_term f with + match EConstr.kind sigma f with | Evar (evk,args) -> (match snd (Evd.evar_source evk sigma) with Evar_kinds.MatchingVar (true,id) -> Some id @@ -153,17 +163,17 @@ let pattern_of_constr env sigma t = | Ind (sp,u) -> PRef (canonical_gr (IndRef sp)) | Construct (sp,u) -> PRef (canonical_gr (ConstructRef sp)) | Proj (p, c) -> - pattern_of_constr env (Retyping.expand_projection env sigma p (EConstr.of_constr c) []) + pattern_of_constr env (EConstr.of_constr (Retyping.expand_projection env sigma p c [])) | Evar (evk,ctxt as ev) -> (match snd (Evd.evar_source evk sigma) with | Evar_kinds.MatchingVar (b,id) -> - let ty = Evarutil.nf_evar sigma (existential_type sigma ev) in + let ty = existential_type sigma ev in let () = ignore (pattern_of_constr env ty) in assert (not b); PMeta (Some id) | Evar_kinds.GoalEvar -> PEvar (evk,Array.map (pattern_of_constr env) ctxt) | _ -> - let ty = Evarutil.nf_evar sigma (existential_type sigma ev) in + let ty = existential_type sigma ev in let () = ignore (pattern_of_constr env ty) in PMeta None) | Case (ci,p,a,br) -> @@ -178,8 +188,13 @@ let pattern_of_constr env sigma t = in PCase (cip, pattern_of_constr env p, pattern_of_constr env a, Array.to_list (Array.mapi branch_of_constr br)) - | Fix f -> PFix f - | CoFix f -> PCoFix f in + | Fix (idx, (nas, cs, ts)) -> + let inj c = EConstr.to_constr sigma c in + PFix (idx, (nas, Array.map inj cs, Array.map inj ts)) + | CoFix (idx, (nas, cs, ts)) -> + let inj c = EConstr.to_constr sigma c in + PCoFix (idx, (nas, Array.map inj cs, Array.map inj ts)) + in pattern_of_constr env t (* To process patterns, we need a translation without typing at all. *) @@ -220,7 +235,7 @@ let instantiate_pattern env sigma lvar c = ctx in let c = substl inst c in - pattern_of_constr env sigma c + pattern_of_constr env sigma (EConstr.of_constr c) with Not_found (* List.index failed *) -> let vars = List.map_filter (function Name id -> Some id | _ -> None) vars in @@ -245,7 +260,7 @@ let rec subst_pattern subst pat = | PRef ref -> let ref',t = subst_global subst ref in if ref' == ref then pat else - pattern_of_constr (Global.env()) Evd.empty t + pattern_of_constr (Global.env()) Evd.empty (EConstr.of_constr t) | PVar _ | PEvar _ | PRel _ -> pat diff --git a/pretyping/patternops.mli b/pretyping/patternops.mli index 1f63565d6..93d2c859a 100644 --- a/pretyping/patternops.mli +++ b/pretyping/patternops.mli @@ -7,6 +7,7 @@ (************************************************************************) open Term +open EConstr open Globnames open Glob_term open Mod_subst @@ -32,7 +33,7 @@ val head_pattern_bound : constr_pattern -> global_reference (** [head_of_constr_reference c] assumes [r] denotes a reference and returns its label; raises an anomaly otherwise *) -val head_of_constr_reference : Term.constr -> global_reference +val head_of_constr_reference : Evd.evar_map -> constr -> global_reference (** [pattern_of_constr c] translates a term [c] with metavariables into a pattern; currently, no destructor (Cases, Fix, Cofix) and no diff --git a/tactics/hints.ml b/tactics/hints.ml index 63d10573a..b2aa02191 100644 --- a/tactics/hints.ml +++ b/tactics/hints.ml @@ -746,7 +746,7 @@ let make_exact_entry env sigma pri poly ?(name=PathAny) (c, cty, ctx) = match kind_of_term cty with | Prod _ -> failwith "make_exact_entry" | _ -> - let pat = Patternops.pattern_of_constr env sigma cty in + let pat = Patternops.pattern_of_constr env sigma (EConstr.of_constr cty) in let hd = try head_pattern_bound pat with BoundPattern -> failwith "make_exact_entry" @@ -767,7 +767,7 @@ let make_apply_entry env sigma (eapply,hnf,verbose) pri poly ?(name=PathAny) (c, let sigma' = Evd.merge_context_set univ_flexible sigma ctx in let ce = mk_clenv_from_env env sigma' None (c,cty) in let c' = clenv_type (* ~reduce:false *) ce in - let pat = Patternops.pattern_of_constr env ce.evd c' in + let pat = Patternops.pattern_of_constr env ce.evd (EConstr.of_constr c') in let hd = try head_pattern_bound pat with BoundPattern -> failwith "make_apply_entry" in @@ -911,11 +911,11 @@ let make_trivial env sigma poly ?(name=PathAny) r = let c,ctx = fresh_global_or_constr env sigma poly r in let sigma = Evd.merge_context_set univ_flexible sigma ctx in let t = hnf_constr env sigma (EConstr.of_constr (unsafe_type_of env sigma (EConstr.of_constr c))) in - let hd = head_of_constr_reference (head_constr sigma t) in + let hd = head_of_constr_reference sigma (EConstr.of_constr (head_constr sigma t)) in let ce = mk_clenv_from_env env sigma None (c,t) in (Some hd, { pri=1; poly = poly; - pat = Some (Patternops.pattern_of_constr env ce.evd (clenv_type ce)); + pat = Some (Patternops.pattern_of_constr env ce.evd (EConstr.of_constr (clenv_type ce))); name = name; db = None; secvars = secvars_of_constr env c; @@ -1013,7 +1013,7 @@ let subst_autohint (subst, obj) = let subst_key gr = let (lab'', elab') = subst_global subst gr in let gr' = - (try head_of_constr_reference (head_constr_bound Evd.empty (** FIXME *) elab') + (try head_of_constr_reference Evd.empty (EConstr.of_constr (head_constr_bound Evd.empty (** FIXME *) elab')) with Bound -> lab'') in if gr' == gr then gr else gr' in -- cgit v1.2.3 From 258c8502eafd3e078a5c7478a452432b5c046f71 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sun, 6 Nov 2016 19:59:28 +0100 Subject: Constr_matching API using EConstr. --- ltac/tactic_matching.ml | 4 +- plugins/quote/quote.ml | 2 +- pretyping/constr_matching.ml | 105 +++++++++++++++++++++++++----------------- pretyping/constr_matching.mli | 3 +- pretyping/tacred.ml | 9 ++-- proofs/tacmach.ml | 6 +-- tactics/auto.ml | 2 +- tactics/class_tactics.ml | 2 +- tactics/hipattern.ml | 4 +- toplevel/search.ml | 10 ++-- 10 files changed, 84 insertions(+), 63 deletions(-) diff --git a/ltac/tactic_matching.ml b/ltac/tactic_matching.ml index 43a58f6b2..58998c96e 100644 --- a/ltac/tactic_matching.ml +++ b/ltac/tactic_matching.ml @@ -233,7 +233,7 @@ module PatternMatching (E:StaticEnvironment) = struct | Term p -> begin try - put_subst (Constr_matching.extended_matches E.env E.sigma p term) <*> + put_subst (Constr_matching.extended_matches E.env E.sigma p (EConstr.of_constr term)) <*> return lhs with Constr_matching.PatternMatchingFailure -> fail end @@ -252,7 +252,7 @@ module PatternMatching (E:StaticEnvironment) = struct | Some nctx -> Proofview.tclOR (k lhs nctx) (fun e -> (map s e).stream k ctx) } in - map (Constr_matching.match_subterm_gen E.env E.sigma with_app_context p term) imatching_error + map (Constr_matching.match_subterm_gen E.env E.sigma with_app_context p (EConstr.of_constr term)) imatching_error (** [rule_match_term term rule] matches the term [term] with the diff --git a/plugins/quote/quote.ml b/plugins/quote/quote.ml index a13948f77..7b6d502b5 100644 --- a/plugins/quote/quote.ml +++ b/plugins/quote/quote.ml @@ -402,7 +402,7 @@ let quote_terms ivs lc = match l with | (lhs, rhs)::tail -> begin try - let s1 = Id.Map.bindings (matches (Global.env ()) Evd.empty rhs c) in + let s1 = Id.Map.bindings (matches (Global.env ()) Evd.empty rhs (EConstr.of_constr c)) in let s2 = List.map (fun (i,c_i) -> (coerce_meta_out i,aux c_i)) s1 in Termops.subst_meta s2 lhs diff --git a/pretyping/constr_matching.ml b/pretyping/constr_matching.ml index 1261844a0..ecf6b1121 100644 --- a/pretyping/constr_matching.ml +++ b/pretyping/constr_matching.ml @@ -57,14 +57,15 @@ let warn_meta_collision = strbrk " and a metavariable of same name.") -let constrain n (ids, m as x) (names, terms as subst) = +let constrain sigma n (ids, m) (names, terms as subst) = + let open EConstr in try let (ids', m') = Id.Map.find n terms in - if List.equal Id.equal ids ids' && eq_constr m m' then subst + if List.equal Id.equal ids ids' && eq_constr sigma m (EConstr.of_constr m') then subst else raise PatternMatchingFailure with Not_found -> let () = if Id.Map.mem n names then warn_meta_collision n in - (names, Id.Map.add n x terms) + (names, Id.Map.add n (ids, EConstr.Unsafe.to_constr m) terms) let add_binders na1 na2 binding_vars (names, terms as subst) = match na1, na2 with @@ -82,8 +83,9 @@ let add_binders na1 na2 binding_vars (names, terms as subst) = let rec build_lambda sigma vars ctx m = match vars with | [] -> let len = List.length ctx in - lift (-1 * len) m + EConstr.Vars.lift (-1 * len) m | n :: vars -> + let open EConstr in (* change [ x1 ... xn y z1 ... zm |- t ] into [ x1 ... xn z1 ... zm |- lam y. t ] *) let len = List.length ctx in @@ -92,7 +94,7 @@ let rec build_lambda sigma vars ctx m = match vars with else if Int.equal i (pred n) then mkRel 1 else mkRel (i + 1) in - let m = substl (List.init len init) m in + let m = Vars.substl (List.init len init) m in let pre, suf = List.chop (pred n) ctx in match suf with | [] -> assert false @@ -100,7 +102,7 @@ let rec build_lambda sigma vars ctx m = match vars with let map i = if i > n then pred i else i in let vars = List.map map vars in (** Check that the abstraction is legal *) - let frels = free_rels sigma (EConstr.of_constr t) in + let frels = free_rels sigma t in let brels = List.fold_right Int.Set.add vars Int.Set.empty in let () = if not (Int.Set.subset frels brels) then raise PatternMatchingFailure in (** Create the abstraction *) @@ -123,41 +125,55 @@ let rec extract_bound_aux k accu frels ctx = match ctx with let extract_bound_vars frels ctx = extract_bound_aux 1 Id.Set.empty frels ctx -let dummy_constr = mkProp +let dummy_constr = EConstr.mkProp let make_renaming ids = function | (Name id, Name _, _) -> begin - try mkRel (List.index Id.equal id ids) + try EConstr.mkRel (List.index Id.equal id ids) with Not_found -> dummy_constr end | _ -> dummy_constr +let local_assum (na, t) = + let inj = EConstr.Unsafe.to_constr in + LocalAssum (na, inj t) + +let local_def (na, b, t) = + let inj = EConstr.Unsafe.to_constr in + LocalDef (na, inj b, inj t) + +let to_fix (idx, (nas, cs, ts)) = + let inj = EConstr.of_constr in + (idx, (nas, Array.map inj cs, Array.map inj ts)) + let merge_binding sigma allow_bound_rels ctx n cT subst = let c = match ctx with | [] -> (* Optimization *) ([], cT) | _ -> - let frels = free_rels sigma (EConstr.of_constr cT) in + let open EConstr in + let frels = free_rels sigma cT in if allow_bound_rels then let vars = extract_bound_vars frels ctx in let ordered_vars = Id.Set.elements vars in let rename binding = make_renaming ordered_vars binding in let renaming = List.map rename ctx in - (ordered_vars, substl renaming cT) + (ordered_vars, Vars.substl renaming cT) else let depth = List.length ctx in let min_elt = try Int.Set.min_elt frels with Not_found -> succ depth in if depth < min_elt then - ([], lift (- depth) cT) + ([], Vars.lift (- depth) cT) else raise PatternMatchingFailure in - constrain n c subst + constrain sigma n c subst let matches_core env sigma convert allow_partial_app allow_bound_rels (binding_vars,pat) c = + let open EConstr in let convref ref c = - match ref, kind_of_term c with + match ref, EConstr.kind sigma c with | VarRef id, Var id' -> Names.id_eq id id' | ConstRef c, Const (c',_) -> Names.eq_constant c c' | IndRef i, Ind (i', _) -> Names.eq_ind i i' @@ -165,12 +181,12 @@ let matches_core env sigma convert allow_partial_app allow_bound_rels | _, _ -> (if convert then let sigma,c' = Evd.fresh_global env sigma ref in - is_conv env sigma (EConstr.of_constr c') (EConstr.of_constr c) + is_conv env sigma (EConstr.of_constr c') c else false) in let rec sorec ctx env subst p t = - let cT = strip_outer_cast sigma (EConstr.of_constr t) in - match p,kind_of_term cT with + let cT = EConstr.of_constr (strip_outer_cast sigma t) in + match p, EConstr.kind sigma cT with | PSoApp (n,args),m -> let fold (ans, seen) = function | PRel n -> @@ -179,9 +195,9 @@ let matches_core env sigma convert allow_partial_app allow_bound_rels | _ -> error "Only bound indices allowed in second order pattern matching." in let relargs, relset = List.fold_left fold ([], Int.Set.empty) args in - let frels = free_rels sigma (EConstr.of_constr cT) in + let frels = free_rels sigma cT in if Int.Set.subset frels relset then - constrain n ([], build_lambda sigma relargs ctx cT) subst + constrain sigma n ([], build_lambda sigma relargs ctx cT) subst else raise PatternMatchingFailure @@ -219,15 +235,15 @@ let matches_core env sigma convert allow_partial_app allow_bound_rels | Some n -> merge_binding sigma allow_bound_rels ctx n c subst in Array.fold_left2 (sorec ctx env) subst args1 args22 else (* Might be a projection on the right *) - match kind_of_term c2 with + match EConstr.kind sigma c2 with | Proj (pr, c) when not (Projection.unfolded pr) -> - (try let term = Retyping.expand_projection env sigma pr (EConstr.of_constr c) (Array.map_to_list EConstr.of_constr args2) in - sorec ctx env subst p term + (try let term = Retyping.expand_projection env sigma pr c (Array.to_list args2) in + sorec ctx env subst p (EConstr.of_constr term) with Retyping.RetypeError _ -> raise PatternMatchingFailure) | _ -> raise PatternMatchingFailure) | PApp (c1,arg1), App (c2,arg2) -> - (match c1, kind_of_term c2 with + (match c1, EConstr.kind sigma c2 with | PRef (ConstRef r), Proj (pr,c) when not (eq_constant r (Projection.constant pr)) || Projection.unfolded pr -> raise PatternMatchingFailure @@ -237,8 +253,8 @@ let matches_core env sigma convert allow_partial_app allow_bound_rels with Invalid_argument _ -> raise PatternMatchingFailure else raise PatternMatchingFailure | _, Proj (pr,c) when not (Projection.unfolded pr) -> - (try let term = Retyping.expand_projection env sigma pr (EConstr.of_constr c) (Array.map_to_list EConstr.of_constr arg2) in - sorec ctx env subst p term + (try let term = Retyping.expand_projection env sigma pr c (Array.to_list arg2) in + sorec ctx env subst p (EConstr.of_constr term) with Retyping.RetypeError _ -> raise PatternMatchingFailure) | _, _ -> try Array.fold_left2 (sorec ctx env) (sorec ctx env subst c1 c2) arg1 arg2 @@ -249,32 +265,32 @@ let matches_core env sigma convert allow_partial_app allow_bound_rels raise PatternMatchingFailure | PApp (c, args), Proj (pr, c2) -> - (try let term = Retyping.expand_projection env sigma pr (EConstr.of_constr c2) [] in - sorec ctx env subst p term + (try let term = Retyping.expand_projection env sigma pr c2 [] in + sorec ctx env subst p (EConstr.of_constr term) with Retyping.RetypeError _ -> raise PatternMatchingFailure) | PProj (p1,c1), Proj (p2,c2) when Projection.equal p1 p2 -> sorec ctx env subst c1 c2 | PProd (na1,c1,d1), Prod(na2,c2,d2) -> - sorec ((na1,na2,c2)::ctx) (Environ.push_rel (LocalAssum (na2,c2)) env) + sorec ((na1,na2,c2)::ctx) (Environ.push_rel (local_assum (na2,c2)) env) (add_binders na1 na2 binding_vars (sorec ctx env subst c1 c2)) d1 d2 | PLambda (na1,c1,d1), Lambda(na2,c2,d2) -> - sorec ((na1,na2,c2)::ctx) (Environ.push_rel (LocalAssum (na2,c2)) env) + sorec ((na1,na2,c2)::ctx) (Environ.push_rel (local_assum (na2,c2)) env) (add_binders na1 na2 binding_vars (sorec ctx env subst c1 c2)) d1 d2 | PLetIn (na1,c1,d1), LetIn(na2,c2,t2,d2) -> - sorec ((na1,na2,t2)::ctx) (Environ.push_rel (LocalDef (na2,c2,t2)) env) + sorec ((na1,na2,t2)::ctx) (Environ.push_rel (local_def (na2,c2,t2)) env) (add_binders na1 na2 binding_vars (sorec ctx env subst c1 c2)) d1 d2 | PIf (a1,b1,b1'), Case (ci,_,a2,[|b2;b2'|]) -> - let ctx_b2,b2 = decompose_lam_n_decls ci.ci_cstr_ndecls.(0) b2 in - let ctx_b2',b2' = decompose_lam_n_decls ci.ci_cstr_ndecls.(1) b2' in + let ctx_b2,b2 = decompose_lam_n_decls sigma ci.ci_cstr_ndecls.(0) b2 in + let ctx_b2',b2' = decompose_lam_n_decls sigma ci.ci_cstr_ndecls.(1) b2' in let n = Context.Rel.length ctx_b2 in let n' = Context.Rel.length ctx_b2' in - if noccur_between 1 n b2 && noccur_between 1 n' b2' then - let f l (LocalAssum (na,t) | LocalDef (na,_,t)) = (Anonymous,na,t)::l in + if Vars.noccur_between sigma 1 n b2 && Vars.noccur_between sigma 1 n' b2' then + let f l (LocalAssum (na,t) | LocalDef (na,_,t)) = (Anonymous,na,EConstr.of_constr t)::l in let ctx_br = List.fold_left f ctx ctx_b2 in let ctx_br' = List.fold_left f ctx ctx_b2' in let b1 = lift_pattern n b1 and b1' = lift_pattern n' b1' in @@ -306,8 +322,8 @@ let matches_core env sigma convert allow_partial_app allow_bound_rels let chk_head = sorec ctx env (sorec ctx env subst a1 a2) p1 p2 in List.fold_left chk_branch chk_head br1 - | PFix c1, Fix _ when eq_constr (mkFix c1) cT -> subst - | PCoFix c1, CoFix _ when eq_constr (mkCoFix c1) cT -> subst + | PFix c1, Fix _ when eq_constr sigma (mkFix (to_fix c1)) cT -> subst + | PCoFix c1, CoFix _ when eq_constr sigma (mkCoFix (to_fix c1)) cT -> subst | _ -> raise PatternMatchingFailure in @@ -328,13 +344,14 @@ type matching_result = { m_sub : bound_ident_map * patvar_map; m_ctx : constr; } -let mkresult s c n = IStream.Cons ( { m_sub=s; m_ctx=c; } , (IStream.thunk n) ) +let mkresult s c n = IStream.Cons ( { m_sub=s; m_ctx=EConstr.Unsafe.to_constr c; } , (IStream.thunk n) ) let isPMeta = function PMeta _ -> true | _ -> false let matches_head env sigma pat c = + let open EConstr in let head = - match pat, kind_of_term c with + match pat, EConstr.kind sigma c with | PApp (c1,arg1), App (c2,arg2) -> if isPMeta c1 then c else let n1 = Array.length arg1 in @@ -345,6 +362,7 @@ let matches_head env sigma pat c = (* Tells if it is an authorized occurrence and if the instance is closed *) let authorized_occ env sigma partial_app closed pat c mk_ctx = + let open EConstr in try let subst = matches_core_closed env sigma false partial_app pat c in if closed && Id.Map.exists (fun _ c -> not (closed0 c)) (snd subst) @@ -356,9 +374,10 @@ let subargs env v = Array.map_to_list (fun c -> (env, c)) v (* Tries to match a subterm of [c] with [pat] *) let sub_match ?(partial_app=false) ?(closed=true) env sigma pat c = + let open EConstr in let rec aux env c mk_ctx next = let here = authorized_occ env sigma partial_app closed pat c mk_ctx in - let next () = match kind_of_term c with + let next () = match EConstr.kind sigma c with | Cast (c1,k,c2) -> let next_mk_ctx = function | [c1] -> mk_ctx (mkCast (c1, k, c2)) @@ -370,21 +389,21 @@ let sub_match ?(partial_app=false) ?(closed=true) env sigma pat c = | [c1; c2] -> mk_ctx (mkLambda (x, c1, c2)) | _ -> assert false in - let env' = Environ.push_rel (LocalAssum (x,c1)) env in + let env' = Environ.push_rel (local_assum (x,c1)) env in try_aux [(env, c1); (env', c2)] next_mk_ctx next | Prod (x,c1,c2) -> let next_mk_ctx = function | [c1; c2] -> mk_ctx (mkProd (x, c1, c2)) | _ -> assert false in - let env' = Environ.push_rel (LocalAssum (x,c1)) env in + let env' = Environ.push_rel (local_assum (x,c1)) env in try_aux [(env, c1); (env', c2)] next_mk_ctx next | LetIn (x,c1,t,c2) -> let next_mk_ctx = function | [c1; c2] -> mk_ctx (mkLetIn (x, c1, t, c2)) | _ -> assert false in - let env' = Environ.push_rel (LocalDef (x,c1,t)) env in + let env' = Environ.push_rel (local_def (x,c1,t)) env in try_aux [(env, c1); (env', c2)] next_mk_ctx next | App (c1,lc) -> let topdown = true in @@ -440,8 +459,8 @@ let sub_match ?(partial_app=false) ?(closed=true) env sigma pat c = let next_mk_ctx le = mk_ctx (mkProj (p,List.hd le)) in if partial_app then try - let term = Retyping.expand_projection env sigma p (EConstr.of_constr c') [] in - aux env term mk_ctx next + let term = Retyping.expand_projection env sigma p c' [] in + aux env (EConstr.of_constr term) mk_ctx next with Retyping.RetypeError _ -> next () else try_aux [env, c'] next_mk_ctx next diff --git a/pretyping/constr_matching.mli b/pretyping/constr_matching.mli index ee6c5141b..32bb48c93 100644 --- a/pretyping/constr_matching.mli +++ b/pretyping/constr_matching.mli @@ -10,6 +10,7 @@ open Names open Term +open EConstr open Environ open Pattern @@ -63,7 +64,7 @@ val matches_conv : env -> Evd.evar_map -> constr_pattern -> constr -> patvar_map (whose hole is denoted here with [special_meta]) *) type matching_result = { m_sub : bound_ident_map * patvar_map; - m_ctx : constr } + m_ctx : Constr.t } (** [match_subterm n pat c] returns the substitution and the context corresponding to each **closed** subterm of [c] matching [pat]. *) diff --git a/pretyping/tacred.ml b/pretyping/tacred.ml index a3983737d..9581db23d 100644 --- a/pretyping/tacred.ml +++ b/pretyping/tacred.ml @@ -962,9 +962,10 @@ let simpl env sigma c = strong whd_simpl env sigma c (* Reduction at specific subterms *) let matches_head env sigma c t = - match kind_of_term t with + let open EConstr in + match EConstr.kind sigma t with | App (f,_) -> Constr_matching.matches env sigma c f - | Proj (p, _) -> Constr_matching.matches env sigma c (mkConst (Projection.constant p)) + | Proj (p, _) -> Constr_matching.matches env sigma c (mkConstU (Projection.constant p, Univ.Instance.empty)) | _ -> raise Constr_matching.PatternMatchingFailure (** FIXME: Specific function to handle projections: it ignores what happens on the @@ -999,8 +1000,8 @@ let e_contextually byhead (occs,c) f = { e_redfun = begin fun env sigma t -> else try let subst = - if byhead then matches_head env sigma c t - else Constr_matching.matches env sigma c t in + if byhead then matches_head env sigma c (EConstr.of_constr t) + else Constr_matching.matches env sigma c (EConstr.of_constr t) in let ok = if nowhere_except_in then Int.List.mem !pos locs else not (Int.List.mem !pos locs) in diff --git a/proofs/tacmach.ml b/proofs/tacmach.ml index 148f9d675..5e9c0ba8e 100644 --- a/proofs/tacmach.ml +++ b/proofs/tacmach.ml @@ -106,8 +106,8 @@ let pf_reduce_to_atomic_ind = pf_reduce reduce_to_atomic_ind let pf_hnf_type_of gls = EConstr.of_constr %> pf_get_type_of gls %> pf_whd_all gls -let pf_is_matching = pf_apply Constr_matching.is_matching_conv -let pf_matches = pf_apply Constr_matching.matches_conv +let pf_is_matching gl p c = pf_apply Constr_matching.is_matching_conv gl p (EConstr.of_constr c) +let pf_matches gl p c = pf_apply Constr_matching.matches_conv gl p (EConstr.of_constr c) (********************************************) (* Definition of the most primitive tactics *) @@ -232,7 +232,7 @@ module New = struct let pf_hnf_type_of gl t = pf_whd_all gl (pf_get_type_of gl (EConstr.of_constr t)) - let pf_matches gl pat t = pf_apply Constr_matching.matches_conv gl pat t + let pf_matches gl pat t = pf_apply Constr_matching.matches_conv gl pat (EConstr.of_constr t) let pf_whd_all gl t = pf_apply whd_all gl t let pf_compute gl t = pf_apply compute gl t diff --git a/tactics/auto.ml b/tactics/auto.ml index 17fe7362d..7462b8d85 100644 --- a/tactics/auto.ml +++ b/tactics/auto.ml @@ -139,7 +139,7 @@ let conclPattern concl pat tac = | None -> Proofview.tclUNIT Id.Map.empty | Some pat -> try - Proofview.tclUNIT (Constr_matching.matches env sigma pat concl) + Proofview.tclUNIT (Constr_matching.matches env sigma pat (EConstr.of_constr concl)) with Constr_matching.PatternMatchingFailure -> Tacticals.New.tclZEROMSG (str "conclPattern") in diff --git a/tactics/class_tactics.ml b/tactics/class_tactics.ml index a31e581e8..bef43d20b 100644 --- a/tactics/class_tactics.ml +++ b/tactics/class_tactics.ml @@ -305,7 +305,7 @@ let matches_pattern concl pat = | None -> Proofview.tclUNIT () | Some pat -> let sigma = Sigma.to_evar_map sigma in - if Constr_matching.is_matching env sigma pat concl then + if Constr_matching.is_matching env sigma pat (EConstr.of_constr concl) then Proofview.tclUNIT () else Tacticals.New.tclZEROMSG (str "conclPattern") diff --git a/tactics/hipattern.ml b/tactics/hipattern.ml index a42a51fc0..d27e4afb7 100644 --- a/tactics/hipattern.ml +++ b/tactics/hipattern.ml @@ -276,8 +276,8 @@ let coq_refl_jm_pattern = open Globnames -let is_matching x y = is_matching (Global.env ()) Evd.empty x y -let matches x y = matches (Global.env ()) Evd.empty x y +let is_matching x y = is_matching (Global.env ()) Evd.empty x (EConstr.of_constr y) +let matches x y = matches (Global.env ()) Evd.empty x (EConstr.of_constr y) let match_with_equation t = if not (isApp t) then raise NoEquationFound; diff --git a/toplevel/search.ml b/toplevel/search.ml index c0bcc403c..653d4ac5c 100644 --- a/toplevel/search.ml +++ b/toplevel/search.ml @@ -113,7 +113,7 @@ let generic_search glnumopt fn = (** FIXME: this is quite dummy, we may find a more efficient algorithm. *) let rec pattern_filter pat ref env typ = let typ = Termops.strip_outer_cast Evd.empty (EConstr.of_constr typ) (** FIXME *) in - if Constr_matching.is_matching env Evd.empty pat typ then true + if Constr_matching.is_matching env Evd.empty pat (EConstr.of_constr typ) then true else match kind_of_term typ with | Prod (_, _, typ) | LetIn (_, _, _, typ) -> pattern_filter pat ref env typ @@ -121,7 +121,7 @@ let rec pattern_filter pat ref env typ = let rec head_filter pat ref env typ = let typ = Termops.strip_outer_cast Evd.empty (EConstr.of_constr typ) (** FIXME *) in - if Constr_matching.is_matching_head env Evd.empty pat typ then true + if Constr_matching.is_matching_head env Evd.empty pat (EConstr.of_constr typ) then true else match kind_of_term typ with | Prod (_, _, typ) | LetIn (_, _, _, typ) -> head_filter pat ref env typ @@ -151,7 +151,7 @@ let name_of_reference ref = Id.to_string (basename_of_global ref) let search_about_filter query gr env typ = match query with | GlobSearchSubPattern pat -> - Constr_matching.is_matching_appsubterm ~closed:false env Evd.empty pat typ + Constr_matching.is_matching_appsubterm ~closed:false env Evd.empty pat (EConstr.of_constr typ) | GlobSearchString s -> String.string_contains ~where:(name_of_reference gr) ~what:s @@ -265,12 +265,12 @@ let interface_search = toggle (Str.string_match regexp id 0) flag in let match_type (pat, flag) = - toggle (Constr_matching.is_matching env Evd.empty pat constr) flag + toggle (Constr_matching.is_matching env Evd.empty pat (EConstr.of_constr constr)) flag in let match_subtype (pat, flag) = toggle (Constr_matching.is_matching_appsubterm ~closed:false - env Evd.empty pat constr) flag + env Evd.empty pat (EConstr.of_constr constr)) flag in let match_module (mdl, flag) = toggle (Libnames.is_dirpath_prefix_of mdl path) flag -- cgit v1.2.3 From b77579ac873975a15978c5a4ecf312d577746d26 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sun, 6 Nov 2016 21:59:18 +0100 Subject: Tacred API using EConstr. --- engine/termops.ml | 25 +++- engine/termops.mli | 5 +- ltac/extratactics.ml4 | 4 +- plugins/funind/recdef.ml | 2 +- plugins/ssrmatching/ssrmatching.ml4 | 5 +- pretyping/evarconv.ml | 5 +- pretyping/find_subterm.ml | 17 ++- pretyping/find_subterm.mli | 3 +- pretyping/tacred.ml | 278 ++++++++++++++++++++---------------- pretyping/tacred.mli | 10 +- pretyping/unification.ml | 6 +- proofs/redexpr.ml | 7 +- tactics/tactics.ml | 4 +- toplevel/himsg.ml | 1 + 14 files changed, 212 insertions(+), 160 deletions(-) diff --git a/engine/termops.ml b/engine/termops.ml index 26b22be4e..f191e2dc1 100644 --- a/engine/termops.ml +++ b/engine/termops.ml @@ -320,8 +320,19 @@ let map_constr_with_named_binders g f l c = match kind_of_term c with time being almost those of the ML representation (except for (co-)fixpoint) *) +let local_assum (na, t) = + let open RelDecl in + let inj = EConstr.Unsafe.to_constr in + LocalAssum (na, inj t) + +let local_def (na, b, t) = + let open RelDecl in + let inj = EConstr.Unsafe.to_constr in + LocalDef (na, inj b, inj t) + let fold_rec_types g (lna,typarray,_) e = - let ctxt = Array.map2_i (fun i na t -> RelDecl.LocalAssum (na, lift i t)) lna typarray in + let open EConstr in + let ctxt = Array.map2_i (fun i na t -> local_assum (na, Vars.lift i t)) lna typarray in Array.fold_left (fun e assum -> g assum e) e ctxt let map_left2 f a g b = @@ -336,9 +347,9 @@ let map_left2 f a g b = r, s end -let map_constr_with_binders_left_to_right g f l c = - let open RelDecl in - match kind_of_term c with +let map_constr_with_binders_left_to_right sigma g f l c = + let open EConstr in + match EConstr.kind sigma c with | (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _ | Construct _) -> c | Cast (b,k,t) -> @@ -348,18 +359,18 @@ let map_constr_with_binders_left_to_right g f l c = else mkCast (b',k,t') | Prod (na,t,b) -> let t' = f l t in - let b' = f (g (LocalAssum (na,t)) l) b in + let b' = f (g (local_assum (na,t)) l) b in if t' == t && b' == b then c else mkProd (na, t', b') | Lambda (na,t,b) -> let t' = f l t in - let b' = f (g (LocalAssum (na,t)) l) b in + let b' = f (g (local_assum (na,t)) l) b in if t' == t && b' == b then c else mkLambda (na, t', b') | LetIn (na,bo,t,b) -> let bo' = f l bo in let t' = f l t in - let b' = f (g (LocalDef (na,bo,t)) l) b in + let b' = f (g (local_def (na,bo,t)) l) b in if bo' == bo && t' == t && b' == b then c else mkLetIn (na, bo', t', b') | App (c,[||]) -> assert false diff --git a/engine/termops.mli b/engine/termops.mli index 24c2c82f2..4becca907 100644 --- a/engine/termops.mli +++ b/engine/termops.mli @@ -65,9 +65,10 @@ val map_constr_with_named_binders : (Name.t -> 'a -> 'a) -> ('a -> constr -> constr) -> 'a -> constr -> constr val map_constr_with_binders_left_to_right : + Evd.evar_map -> (Context.Rel.Declaration.t -> 'a -> 'a) -> - ('a -> constr -> constr) -> - 'a -> constr -> constr + ('a -> EConstr.constr -> EConstr.constr) -> + 'a -> EConstr.constr -> EConstr.constr val map_constr_with_full_binders : Evd.evar_map -> (Context.Rel.Declaration.t -> 'a -> 'a) -> diff --git a/ltac/extratactics.ml4 b/ltac/extratactics.ml4 index ca1f6e638..e1b468197 100644 --- a/ltac/extratactics.ml4 +++ b/ltac/extratactics.ml4 @@ -300,7 +300,7 @@ let project_hint pri l2r r = let sigma, c = Evd.fresh_global env sigma gr in let t = Retyping.get_type_of env sigma (EConstr.of_constr c) in let t = - Tacred.reduce_to_quantified_ref env sigma (Lazy.force coq_iff_ref) t in + Tacred.reduce_to_quantified_ref env sigma (Lazy.force coq_iff_ref) (EConstr.of_constr t) in let sign,ccl = decompose_prod_assum t in let (a,b) = match snd (decompose_app ccl) with | [a;b] -> (a,b) @@ -738,7 +738,7 @@ let mkCaseEq a : unit Proofview.tactic = let concl = Proofview.Goal.concl gl in let env = Proofview.Goal.env gl in (** FIXME: this looks really wrong. Does anybody really use this tactic? *) - let Sigma (c, _, _) = (Tacred.pattern_occs [Locus.OnlyOccurrences [1], a]).Reductionops.e_redfun env (Sigma.Unsafe.of_evar_map Evd.empty) (EConstr.of_constr concl) in + let Sigma (c, _, _) = (Tacred.pattern_occs [Locus.OnlyOccurrences [1], EConstr.of_constr a]).Reductionops.e_redfun env (Sigma.Unsafe.of_evar_map Evd.empty) (EConstr.of_constr concl) in change_concl c end }; simplest_case a] diff --git a/plugins/funind/recdef.ml b/plugins/funind/recdef.ml index 12ed758ba..bdbf0242d 100644 --- a/plugins/funind/recdef.ml +++ b/plugins/funind/recdef.ml @@ -692,7 +692,7 @@ let mkDestructEq : [Proofview.V82.of_tactic (generalize new_hyps); (fun g2 -> let changefun patvars = { run = fun sigma -> - let redfun = pattern_occs [Locus.AllOccurrencesBut [1], expr] in + let redfun = pattern_occs [Locus.AllOccurrencesBut [1], EConstr.of_constr expr] in redfun.Reductionops.e_redfun (pf_env g2) sigma (EConstr.of_constr (pf_concl g2)) } in Proofview.V82.of_tactic (change_in_concl None changefun) g2); diff --git a/plugins/ssrmatching/ssrmatching.ml4 b/plugins/ssrmatching/ssrmatching.ml4 index e0d99d453..18aeca6fa 100644 --- a/plugins/ssrmatching/ssrmatching.ml4 +++ b/plugins/ssrmatching/ssrmatching.ml4 @@ -836,7 +836,10 @@ let rec uniquize = function | Context.Rel.Declaration.LocalDef (x,_,y) -> Context.Rel.Declaration.LocalAssum(x,y) in Environ.push_rel ctx_item env, h' + 1 in - let f' = map_constr_with_binders_left_to_right inc_h subst_loop acc f in + let self acc c = EConstr.of_constr (subst_loop acc (EConstr.Unsafe.to_constr c)) in + let f = EConstr.of_constr f in + let f' = map_constr_with_binders_left_to_right sigma inc_h self acc f in + let f' = EConstr.Unsafe.to_constr f' in mkApp (f', Array.map_left (subst_loop acc) a) in subst_loop (env,h) c) : find_P), ((fun () -> diff --git a/pretyping/evarconv.ml b/pretyping/evarconv.ml index c8dcb19b4..cdcb993b5 100644 --- a/pretyping/evarconv.ml +++ b/pretyping/evarconv.ml @@ -986,10 +986,9 @@ let apply_on_subterm env evdref f c t = let g decl a = if is_local_assum decl then applyrec acc a else a in mkEvar (evk, Array.of_list (List.map2 g ctx (Array.to_list args))) | _ -> - let self acc c = EConstr.Unsafe.to_constr (applyrec acc (EConstr.of_constr c)) in - EConstr.of_constr (map_constr_with_binders_left_to_right + map_constr_with_binders_left_to_right !evdref (fun d (env,(k,c)) -> (push_rel d env, (k+1,Vars.lift 1 c))) - self acc (EConstr.Unsafe.to_constr t)) + applyrec acc t in applyrec (env,(0,c)) t diff --git a/pretyping/find_subterm.ml b/pretyping/find_subterm.ml index d7f2d54aa..2b243d5b9 100644 --- a/pretyping/find_subterm.ml +++ b/pretyping/find_subterm.ml @@ -99,7 +99,7 @@ type 'a testing_function = { (b,l), b=true means no occurrence except the ones in l and b=false, means all occurrences except the ones in l *) -let replace_term_occ_gen_modulo occs like_first test bywhat cl occ t = +let replace_term_occ_gen_modulo sigma occs like_first test bywhat cl occ t = let (nowhere_except_in,locs) = Locusops.convert_occs occs in let maxocc = List.fold_right max locs 0 in let pos = ref occ in @@ -133,24 +133,23 @@ let replace_term_occ_gen_modulo occs like_first test bywhat cl occ t = with NotUnifiable _ -> subst_below k t and subst_below k t = - let substrec i c = EConstr.Unsafe.to_constr (substrec i (EConstr.of_constr c)) in - EConstr.of_constr (map_constr_with_binders_left_to_right (fun d k -> k+1) substrec k (EConstr.Unsafe.to_constr t)) + map_constr_with_binders_left_to_right sigma (fun d k -> k+1) substrec k t in let t' = substrec 0 t in (!pos, t') -let replace_term_occ_modulo occs test bywhat t = +let replace_term_occ_modulo evd occs test bywhat t = let occs',like_first = match occs with AtOccs occs -> occs,false | LikeFirst -> AllOccurrences,true in EConstr.Unsafe.to_constr (proceed_with_occurrences - (replace_term_occ_gen_modulo occs' like_first test bywhat None) occs' t) + (replace_term_occ_gen_modulo evd occs' like_first test bywhat None) occs' t) -let replace_term_occ_decl_modulo occs test bywhat d = +let replace_term_occ_decl_modulo evd occs test bywhat d = let (plocs,hyploc),like_first = match occs with AtOccs occs -> occs,false | LikeFirst -> (AllOccurrences,InHyp),true in proceed_with_occurrences (map_named_declaration_with_hyploc - (replace_term_occ_gen_modulo plocs like_first test bywhat) + (replace_term_occ_gen_modulo evd plocs like_first test bywhat) hyploc) plocs d @@ -172,7 +171,7 @@ let make_eq_univs_test env evd c = let subst_closed_term_occ env evd occs c t = let test = make_eq_univs_test env evd c in let bywhat () = mkRel 1 in - let t' = replace_term_occ_modulo occs test bywhat t in + let t' = replace_term_occ_modulo evd occs test bywhat t in t', test.testing_state let subst_closed_term_occ_decl env evd occs c d = @@ -182,6 +181,6 @@ let subst_closed_term_occ_decl env evd occs c d = let bywhat () = mkRel 1 in proceed_with_occurrences (map_named_declaration_with_hyploc - (fun _ -> replace_term_occ_gen_modulo plocs like_first test bywhat None) + (fun _ -> replace_term_occ_gen_modulo evd plocs like_first test bywhat None) hyploc) plocs d, test.testing_state diff --git a/pretyping/find_subterm.mli b/pretyping/find_subterm.mli index 49a5dd7f2..e7f0da93f 100644 --- a/pretyping/find_subterm.mli +++ b/pretyping/find_subterm.mli @@ -41,12 +41,13 @@ val make_eq_univs_test : env -> evar_map -> EConstr.constr -> evar_map testing_f matching subterms at the indicated occurrences [occl] with [mk ()]; it turns a NotUnifiable exception raised by the testing function into a SubtermUnificationError. *) -val replace_term_occ_modulo : occurrences or_like_first -> +val replace_term_occ_modulo : evar_map -> occurrences or_like_first -> 'a testing_function -> (unit -> EConstr.constr) -> EConstr.constr -> constr (** [replace_term_occ_decl_modulo] is similar to [replace_term_occ_modulo] but for a named_declaration. *) val replace_term_occ_decl_modulo : + evar_map -> (occurrences * hyp_location_flag) or_like_first -> 'a testing_function -> (unit -> EConstr.constr) -> Context.Named.Declaration.t -> Context.Named.Declaration.t diff --git a/pretyping/tacred.ml b/pretyping/tacred.ml index 9581db23d..9997976c4 100644 --- a/pretyping/tacred.ml +++ b/pretyping/tacred.ml @@ -31,7 +31,7 @@ module NamedDecl = Context.Named.Declaration (* Errors *) type reduction_tactic_error = - InvalidAbstraction of env * Evd.evar_map * constr * (env * Type_errors.type_error) + InvalidAbstraction of env * Evd.evar_map * EConstr.constr * (env * Type_errors.type_error) exception ReductionTacticError of reduction_tactic_error @@ -77,14 +77,14 @@ type evaluable_reference = | EvalConst of constant | EvalVar of Id.t | EvalRel of int - | EvalEvar of existential + | EvalEvar of EConstr.existential -let evaluable_reference_eq r1 r2 = match r1, r2 with +let evaluable_reference_eq sigma r1 r2 = match r1, r2 with | EvalConst c1, EvalConst c2 -> eq_constant c1 c2 | EvalVar id1, EvalVar id2 -> Id.equal id1 id2 | EvalRel i1, EvalRel i2 -> Int.equal i1 i2 | EvalEvar (e1, ctx1), EvalEvar (e2, ctx2) -> - Evar.equal e1 e2 && Array.equal eq_constr ctx1 ctx2 + Evar.equal e1 e2 && Array.equal (EConstr.eq_constr sigma) ctx1 ctx2 | _ -> false let mkEvalRef ref u = @@ -93,15 +93,15 @@ let mkEvalRef ref u = | EvalConst cst -> mkConstU (cst,u) | EvalVar id -> mkVar id | EvalRel n -> mkRel n - | EvalEvar ev -> EConstr.of_constr (Constr.mkEvar ev) + | EvalEvar ev -> EConstr.mkEvar ev -let isEvalRef env c = match kind_of_term c with +let isEvalRef env sigma c = match EConstr.kind sigma c with | Const (sp,_) -> is_evaluable env (EvalConstRef sp) | Var id -> is_evaluable env (EvalVarRef id) | Rel _ | Evar _ -> true | _ -> false -let destEvalRefU c = match kind_of_term c with +let destEvalRefU sigma c = match EConstr.kind sigma c with | Const (cst,u) -> EvalConst cst, u | Var id -> (EvalVar id, Univ.Instance.empty) | Rel n -> (EvalRel n, Univ.Instance.empty) @@ -109,31 +109,39 @@ let destEvalRefU c = match kind_of_term c with | _ -> anomaly (Pp.str "Not an unfoldable reference") let unsafe_reference_opt_value env sigma eval = + let open EConstr in match eval with | EvalConst cst -> (match (lookup_constant cst env).Declarations.const_body with - | Declarations.Def c -> Some (Mod_subst.force_constr c) + | Declarations.Def c -> Some (EConstr.of_constr (Mod_subst.force_constr c)) | _ -> None) | EvalVar id -> - env |> lookup_named id |> NamedDecl.get_value + env |> lookup_named id |> NamedDecl.get_value |> Option.map EConstr.of_constr | EvalRel n -> - env |> lookup_rel n |> RelDecl.map_value (lift n) |> RelDecl.get_value - | EvalEvar ev -> Evd.existential_opt_value sigma ev + env |> lookup_rel n |> RelDecl.map_value (lift n) |> RelDecl.get_value |> Option.map EConstr.of_constr + | EvalEvar ev -> + match EConstr.kind sigma (mkEvar ev) with + | Evar _ -> None + | c -> Some (EConstr.of_kind c) let reference_opt_value env sigma eval u = + let open EConstr in match eval with - | EvalConst cst -> constant_opt_value_in env (cst,u) + | EvalConst cst -> Option.map EConstr.of_constr (constant_opt_value_in env (cst,u)) | EvalVar id -> - env |> lookup_named id |> NamedDecl.get_value + env |> lookup_named id |> NamedDecl.get_value |> Option.map EConstr.of_constr | EvalRel n -> - env |> lookup_rel n |> RelDecl.map_value (lift n) |> RelDecl.get_value - | EvalEvar ev -> Evd.existential_opt_value sigma ev + env |> lookup_rel n |> RelDecl.map_value (lift n) |> RelDecl.get_value |> Option.map EConstr.of_constr + | EvalEvar ev -> + match EConstr.kind sigma (mkEvar ev) with + | Evar _ -> None + | c -> Some (EConstr.of_kind c) exception NotEvaluable let reference_value env sigma c u = match reference_opt_value env sigma c u with | None -> raise NotEvaluable - | Some d -> EConstr.of_constr d + | Some d -> d (************************************************************************) (* Reduction of constants hiding a fixpoint (e.g. for "simpl" tactic). *) @@ -179,6 +187,7 @@ let eval_table = Summary.ref (Cmap.empty : frozen) ~name:"evaluation" *) let check_fix_reversibility sigma labs args ((lv,i),(_,tys,bds)) = + let open EConstr in let n = List.length labs in let nargs = List.length args in if nargs > n then raise Elimconst; @@ -188,8 +197,8 @@ let check_fix_reversibility sigma labs args ((lv,i),(_,tys,bds)) = (function d -> match EConstr.kind sigma d with | Rel k -> if - Array.for_all (noccurn k) tys - && Array.for_all (noccurn (k+nbfix)) bds + Array.for_all (Vars.noccurn sigma k) tys + && Array.for_all (Vars.noccurn sigma (k+nbfix)) bds && k <= n then (k, List.nth labs (k-1)) @@ -223,6 +232,7 @@ let check_fix_reversibility sigma labs args ((lv,i),(_,tys,bds)) = let invert_name labs l na0 env sigma ref = function | Name id -> + let open EConstr in let minfxargs = List.length l in begin match na0 with | Name id' when Id.equal id' id -> @@ -239,12 +249,13 @@ let invert_name labs l na0 env sigma ref = function try match unsafe_reference_opt_value env sigma ref with | None -> None | Some c -> - let labs',ccl = decompose_lam c in - let _, l' = whd_betalet_stack sigma (EConstr.of_constr ccl) in + let labs',ccl = decompose_lam sigma c in + let _, l' = whd_betalet_stack sigma ccl in let labs' = List.map snd labs' in (** ppedrot: there used to be generic equality on terms here *) + let eq_constr c1 c2 = EConstr.eq_constr sigma c1 c2 in if List.equal eq_constr labs' labs && - List.equal (fun c1 c2 -> EConstr.eq_constr sigma c1 c2) l l' then Some (minfxargs,ref) + List.equal eq_constr l l' then Some (minfxargs,ref) else None with Not_found (* Undefined ref *) -> None end @@ -254,20 +265,29 @@ let invert_name labs l na0 env sigma ref = function [compute_consteval_mutual_fix] only one by one, until finding the last one before the Fix if the latter is mutually defined *) +let local_assum (na, t) = + let open Context.Rel.Declaration in + let inj = EConstr.Unsafe.to_constr in + LocalAssum (na, inj t) + +let local_def (na, b, t) = + let open Context.Rel.Declaration in + let inj = EConstr.Unsafe.to_constr in + LocalDef (na, inj b, inj t) + let compute_consteval_direct env sigma ref = + let open EConstr in let rec srec env n labs onlyproj c = - let c',l = whd_betadeltazeta_stack env sigma (EConstr.of_constr c) in - let c' = EConstr.Unsafe.to_constr c' in - match kind_of_term c' with + let c',l = whd_betadeltazeta_stack env sigma c in + match EConstr.kind sigma c' with | Lambda (id,t,g) when List.is_empty l && not onlyproj -> - let open Context.Rel.Declaration in - srec (push_rel (LocalAssum (id,t)) env) (n+1) (EConstr.of_constr t::labs) onlyproj g + srec (push_rel (local_assum (id,t)) env) (n+1) (t::labs) onlyproj g | Fix fix when not onlyproj -> (try check_fix_reversibility sigma labs l fix with Elimconst -> NotAnElimination) - | Case (_,_,d,_) when isRel d && not onlyproj -> EliminationCases n + | Case (_,_,d,_) when isRel sigma d && not onlyproj -> EliminationCases n | Case (_,_,d,_) -> srec env n labs true d - | Proj (p, d) when isRel d -> EliminationProj n + | Proj (p, d) when isRel sigma d -> EliminationProj n | _ -> NotAnElimination in match unsafe_reference_opt_value env sigma ref with @@ -275,14 +295,13 @@ let compute_consteval_direct env sigma ref = | Some c -> srec env 0 [] false c let compute_consteval_mutual_fix env sigma ref = + let open EConstr in let rec srec env minarg labs ref c = - let c',l = whd_betalet_stack sigma (EConstr.of_constr c) in + let c',l = whd_betalet_stack sigma c in let nargs = List.length l in - let c' = EConstr.Unsafe.to_constr c' in - match kind_of_term c' with + match EConstr.kind sigma c' with | Lambda (na,t,g) when List.is_empty l -> - let open Context.Rel.Declaration in - srec (push_rel (LocalAssum (na,t)) env) (minarg+1) (t::labs) ref g + srec (push_rel (local_assum (na,t)) env) (minarg+1) (t::labs) ref g | Fix ((lv,i),(names,_,_)) -> (* Last known constant wrapping Fix is ref = [labs](Fix l) *) (match compute_consteval_direct env sigma ref with @@ -295,9 +314,9 @@ let compute_consteval_mutual_fix env sigma ref = let new_minarg = max (minarg'+minarg-nargs) minarg' in EliminationMutualFix (new_minarg,ref,(refs,infos)) | _ -> assert false) - | _ when isEvalRef env c' -> + | _ when isEvalRef env sigma c' -> (* Forget all \'s and args and do as if we had started with c' *) - let ref,_ = destEvalRefU c' in + let ref,_ = destEvalRefU sigma c' in (match unsafe_reference_opt_value env sigma ref with | None -> anomaly (Pp.str "Should have been trapped by compute_direct") | Some c -> srec env (minarg-nargs) [] ref c) @@ -417,23 +436,25 @@ let solve_arity_problem env sigma fxminargs c = let rec check strict c = let c' = EConstr.of_constr (whd_betaiotazeta sigma c) in let (h,rcargs) = decompose_app_vect sigma c' in - match kind_of_term h with + let rcargs = Array.map EConstr.of_constr rcargs in + let h = EConstr.of_constr h in + match EConstr.kind sigma h with Evar(i,_) when Evar.Map.mem i fxminargs && not (Evd.is_defined !evm i) -> let minargs = Evar.Map.find i fxminargs in if Array.length rcargs < minargs then if strict then set_fix i else raise Partial; - Array.iter (EConstr.of_constr %> check strict) rcargs - | (Var _|Const _) when isEvalRef env h -> - (let ev, u = destEvalRefU h in + Array.iter (check strict) rcargs + | (Var _|Const _) when isEvalRef env sigma h -> + (let ev, u = destEvalRefU sigma h in match reference_opt_value env sigma ev u with | Some h' -> let bak = !evm in - (try Array.iter (EConstr.of_constr %> check false) rcargs + (try Array.iter (check false) rcargs with Partial -> evm := bak; - check strict (EConstr.of_constr (Constr.mkApp(h',rcargs)))) - | None -> Array.iter (EConstr.of_constr %> check strict) rcargs) + check strict (mkApp(h',rcargs))) + | None -> Array.iter (check strict) rcargs) | _ -> EConstr.iter sigma (check strict) c' in check true c; !evm @@ -445,14 +466,16 @@ let substl_checking_arity env subst sigma c = let sigma' = solve_arity_problem env sigma minargs body in (* we propagate the constraints: solved problems are substituted; the other ones are replaced by the function symbol *) - let rec nf_fix c = - match kind_of_term c with - Evar(i,[|fx;f|] as ev) when Evar.Map.mem i minargs -> - (match Evd.existential_opt_value sigma' ev with - Some c' -> c' - | None -> f) - | _ -> map_constr nf_fix c in - EConstr.of_constr (nf_fix (EConstr.Unsafe.to_constr body)) + let rec nf_fix c = match EConstr.kind sigma c with + | Evar (i,[|fx;f|]) when Evar.Map.mem i minargs -> + (** FIXME: find a less hackish way of doing this *) + begin match EConstr.kind sigma' c with + | Evar _ -> f + | c -> EConstr.of_kind c + end + | _ -> EConstr.map sigma nf_fix c + in + nf_fix body type fix_reduction_result = NotReducible | Reduced of (EConstr.t * EConstr.t list) @@ -540,21 +563,21 @@ let match_eval_ref env sigma constr = Some (EvalConst sp, u) | Var id when is_evaluable env (EvalVarRef id) -> Some (EvalVar id, Univ.Instance.empty) | Rel i -> Some (EvalRel i, Univ.Instance.empty) - | Evar (evk, args) -> Some (EvalEvar (evk, Array.map EConstr.Unsafe.to_constr args), Univ.Instance.empty) + | Evar ev -> Some (EvalEvar ev, Univ.Instance.empty) | _ -> None let match_eval_ref_value env sigma constr = - match kind_of_term constr with + match EConstr.kind sigma constr with | Const (sp, u) when is_evaluable env (EvalConstRef sp) -> - Some (constant_value_in env (sp, u)) + Some (EConstr.of_constr (constant_value_in env (sp, u))) | Var id when is_evaluable env (EvalVarRef id) -> - env |> lookup_named id |> NamedDecl.get_value + env |> lookup_named id |> NamedDecl.get_value |> Option.map EConstr.of_constr | Rel n -> - env |> lookup_rel n |> RelDecl.map_value (lift n) |> RelDecl.get_value - | Evar ev -> Evd.existential_opt_value sigma ev + env |> lookup_rel n |> RelDecl.map_value (lift n) |> RelDecl.get_value |> Option.map EConstr.of_constr | _ -> None let special_red_case env sigma whfun (ci, p, c, lf) = + let open EConstr in let rec redrec s = let (constr, cargs) = whfun s in match match_eval_ref env sigma constr with @@ -562,13 +585,12 @@ let special_red_case env sigma whfun (ci, p, c, lf) = (match reference_opt_value env sigma ref u with | None -> raise Redelimination | Some gvalue -> - let gvalue = EConstr.of_constr gvalue in if reducible_mind_case sigma gvalue then reduce_mind_case_use_function constr env sigma {mP=p; mconstr=gvalue; mcargs=cargs; mci=ci; mlf=lf} else - redrec (EConstr.applist(gvalue, cargs))) + redrec (applist(gvalue, cargs))) | None -> if reducible_mind_case sigma constr then reduce_mind_case sigma @@ -688,11 +710,11 @@ let rec red_elim_const env sigma ref u largs = | EliminationMutualFix (min,refgoal,refinfos) when nargs >= min -> let rec descend (ref,u) args = let c = reference_value env sigma ref u in - if evaluable_reference_eq ref refgoal then + if evaluable_reference_eq sigma ref refgoal then (c,args) else let c', lrest = whd_betalet_stack sigma (applist(c,args)) in - descend (destEvalRefU (EConstr.Unsafe.to_constr c')) lrest in + descend (destEvalRefU sigma c') lrest in let (_, midargs as s) = descend (ref,u) largs in let d, lrest = whd_nothing_for_iota env sigma (applist s) in let f = make_elim_fun refinfos u midargs in @@ -803,7 +825,7 @@ and whd_construct_stack env sigma s = | Some (ref, u) -> (match reference_opt_value env sigma ref u with | None -> raise Redelimination - | Some gvalue -> whd_construct_stack env sigma (applist(EConstr.of_constr gvalue, cargs))) + | Some gvalue -> whd_construct_stack env sigma (applist(gvalue, cargs))) | _ -> raise Redelimination (************************************************************************) @@ -816,7 +838,6 @@ and whd_construct_stack env sigma s = let try_red_product env sigma c = let simpfun c = EConstr.of_constr (clos_norm_flags betaiotazeta env sigma c) in - let inj = EConstr.Unsafe.to_constr in let open EConstr in let rec redrec env x = let x = EConstr.of_constr (whd_betaiota sigma x) in @@ -834,8 +855,7 @@ let try_red_product env sigma c = | _ -> simpfun (mkApp (redrec env f, l))) | Cast (c,_,_) -> redrec env c | Prod (x,a,b) -> - let open Context.Rel.Declaration in - mkProd (x, a, redrec (push_rel (LocalAssum (x, inj a)) env) b) + mkProd (x, a, redrec (push_rel (local_assum (x, a)) env) b) | LetIn (x,a,b,t) -> redrec env (Vars.subst1 a t) | Case (ci,p,d,lf) -> simpfun (mkCase (ci,p,redrec env d,lf)) | Proj (p, c) -> @@ -855,7 +875,7 @@ let try_red_product env sigma c = (* to get true one-step reductions *) (match reference_opt_value env sigma ref u with | None -> raise Redelimination - | Some c -> EConstr.of_constr c) + | Some c -> c) | _ -> raise Redelimination) in EConstr.Unsafe.to_constr (redrec env c) @@ -931,9 +951,9 @@ let whd_simpl_orelse_delta_but_fix env sigma c = let open EConstr in let rec redrec s = let (constr, stack as s') = whd_simpl_stack env sigma s in - match match_eval_ref_value env sigma (EConstr.Unsafe.to_constr constr) with + match match_eval_ref_value env sigma constr with | Some c -> - (match kind_of_term (strip_lam c) with + (match EConstr.kind sigma (snd (decompose_lam sigma c)) with | CoFix _ | Fix _ -> s' | Proj (p,t) when (match EConstr.kind sigma constr with @@ -943,8 +963,8 @@ let whd_simpl_orelse_delta_but_fix env sigma c = if List.length stack <= pb.Declarations.proj_npars then (** Do not show the eta-expanded form *) s' - else redrec (applist (EConstr.of_constr c, stack)) - | _ -> redrec (applist(EConstr.of_constr c, stack))) + else redrec (applist (c, stack)) + | _ -> redrec (applist(c, stack))) | None -> s' in let simpfun = clos_norm_flags betaiota env sigma in @@ -973,22 +993,25 @@ let matches_head env sigma c t = of the projection and its eta expanded form. *) let change_map_constr_with_binders_left_to_right g f (env, l as acc) sigma c = - match kind_of_term c with + let open EConstr in + match EConstr.kind sigma c with | Proj (p, r) -> (* Treat specially for partial applications *) - let t = Retyping.expand_projection env sigma p (EConstr.of_constr r) [] in - let hdf, al = destApp t in + let t = Retyping.expand_projection env sigma p r [] in + let t = EConstr.of_constr t in + let hdf, al = destApp sigma t in let a = al.(Array.length al - 1) in let app = (mkApp (hdf, Array.sub al 0 (Array.length al - 1))) in let app' = f acc app in let a' = f acc a in - (match kind_of_term app' with + (match EConstr.kind sigma app' with | App (hdf', al') when hdf' == hdf -> (* Still the same projection, we ignore the change in parameters *) mkProj (p, a') | _ -> mkApp (app', [| a' |])) - | _ -> map_constr_with_binders_left_to_right g f acc c + | _ -> map_constr_with_binders_left_to_right sigma g f acc c let e_contextually byhead (occs,c) f = { e_redfun = begin fun env sigma t -> + let open EConstr in let (nowhere_except_in,locs) = Locusops.convert_occs occs in let maxocc = List.fold_right max locs 0 in let pos = ref 1 in @@ -1000,8 +1023,8 @@ let e_contextually byhead (occs,c) f = { e_redfun = begin fun env sigma t -> else try let subst = - if byhead then matches_head env sigma c (EConstr.of_constr t) - else Constr_matching.matches env sigma c (EConstr.of_constr t) in + if byhead then matches_head env sigma c t + else Constr_matching.matches env sigma c t in let ok = if nowhere_except_in then Int.List.mem !pos locs else not (Int.List.mem !pos locs) in @@ -1012,8 +1035,8 @@ let e_contextually byhead (occs,c) f = { e_redfun = begin fun env sigma t -> (* Skip inner occurrences for stable counting of occurrences *) if locs != [] then ignore (traverse_below (Some (!pos-1)) envc t); - let Sigma (t, evm, _) = (f subst).e_redfun env (Sigma.Unsafe.of_evar_map !evd) (EConstr.of_constr t) in - (evd := Sigma.to_evar_map evm; t) + let Sigma (t, evm, _) = (f subst).e_redfun env (Sigma.Unsafe.of_evar_map !evd) t in + (evd := Sigma.to_evar_map evm; EConstr.of_constr t) end else traverse_below nested envc t @@ -1022,7 +1045,7 @@ let e_contextually byhead (occs,c) f = { e_redfun = begin fun env sigma t -> and traverse_below nested envc t = (* when byhead, find other occurrences without matching again partial application with same head *) - match kind_of_term t with + match EConstr.kind !evd t with | App (f,l) when byhead -> mkApp (f, Array.map_left (traverse nested envc) l) | Proj (p,c) when byhead -> mkProj (p,traverse nested envc c) | _ -> @@ -1030,9 +1053,9 @@ let e_contextually byhead (occs,c) f = { e_redfun = begin fun env sigma t -> (fun d (env,c) -> (push_rel d env,lift_pattern 1 c)) (traverse nested) envc sigma t in - let t' = traverse None (env,c) (EConstr.Unsafe.to_constr t) in + let t' = traverse None (env,c) t in if List.exists (fun o -> o >= !pos) locs then error_invalid_occurrence locs; - Sigma.Unsafe.of_pair (t', !evd) + Sigma.Unsafe.of_pair (EConstr.Unsafe.to_constr t', !evd) end } let contextually byhead occs f env sigma t = @@ -1068,10 +1091,9 @@ let substlin env sigma evalref n (nowhere_except_in,locs) c = incr pos; if ok then value u else c | None -> - let self () c = EConstr.Unsafe.to_constr (substrec () (EConstr.of_constr c)) in - EConstr.of_constr (map_constr_with_binders_left_to_right + map_constr_with_binders_left_to_right sigma (fun _ () -> ()) - self () (EConstr.Unsafe.to_constr c)) + substrec () c in let t' = substrec () c in (!pos, t') @@ -1082,9 +1104,9 @@ let string_of_evaluable_ref env = function string_of_qualid (Nametab.shortest_qualid_of_global (vars_of_env env) (ConstRef kn)) -let unfold env sigma name = +let unfold env sigma name c = if is_evaluable env name then - clos_norm_flags (unfold_red name) env sigma + EConstr.of_constr (clos_norm_flags (unfold_red name) env sigma c) else error (string_of_evaluable_ref env name^" is opaque.") @@ -1102,37 +1124,40 @@ let unfoldoccs env sigma (occs,name) c = | [] -> () | _ -> error_invalid_occurrence rest in - nf_betaiotazeta sigma uc + EConstr.of_constr (nf_betaiotazeta sigma uc) in match occs with - | NoOccurrences -> EConstr.Unsafe.to_constr c + | NoOccurrences -> c | AllOccurrences -> unfold env sigma name c | OnlyOccurrences l -> unfo true l | AllOccurrencesBut l -> unfo false l (* Unfold reduction tactic: *) let unfoldn loccname env sigma c = - EConstr.Unsafe.to_constr (List.fold_left (fun c occname -> EConstr.of_constr (unfoldoccs env sigma occname c)) c loccname) + EConstr.Unsafe.to_constr (List.fold_left (fun c occname -> unfoldoccs env sigma occname c) c loccname) (* Re-folding constants tactics: refold com in term c *) let fold_one_com com env sigma c = + let open EConstr in let rcom = - try red_product env sigma (EConstr.of_constr com) + try EConstr.of_constr (red_product env sigma com) with Redelimination -> error "Not reducible." in (* Reason first on the beta-iota-zeta normal form of the constant as unfold produces it, so that the "unfold f; fold f" configuration works to refold fix expressions *) - let a = subst_term sigma (EConstr.of_constr (clos_norm_flags unfold_side_red env sigma (EConstr.of_constr rcom))) c in - if not (eq_constr a (EConstr.Unsafe.to_constr c)) then - subst1 com a + let a = subst_term sigma (EConstr.of_constr (clos_norm_flags unfold_side_red env sigma rcom)) c in + let a = EConstr.of_constr a in + if not (EConstr.eq_constr sigma a c) then + Vars.subst1 com a else (* Then reason on the non beta-iota-zeta form for compatibility - even if it is probably a useless configuration *) - let a = subst_term sigma (EConstr.of_constr rcom) c in - subst1 com a + let a = subst_term sigma rcom c in + let a = EConstr.of_constr a in + Vars.subst1 com a let fold_commands cl env sigma c = - EConstr.Unsafe.to_constr (List.fold_right (fun com c -> EConstr.of_constr (fold_one_com com env sigma c)) (List.rev cl) c) + EConstr.Unsafe.to_constr (List.fold_right (fun com c -> fold_one_com com env sigma c) (List.rev cl) c) (* call by value reduction functions *) @@ -1150,23 +1175,26 @@ let compute = cbv_betadeltaiota (* gives [na:ta]c' such that c converts to ([na:ta]c' a), abstracting only * the specified occurrences. *) -let abstract_scheme env (locc,a) (c, sigma) = - let a = EConstr.of_constr a in +let abstract_scheme env sigma (locc,a) (c, sigma) = + let open EConstr in let ta = Retyping.get_type_of env sigma a in - let na = named_hd env ta Anonymous in - if occur_meta sigma (EConstr.of_constr ta) then error "Cannot find a type for the generalisation."; + let ta = EConstr.of_constr ta in + let na = named_hd env (EConstr.to_constr sigma ta) Anonymous in + if occur_meta sigma ta then error "Cannot find a type for the generalisation."; if occur_meta sigma a then mkLambda (na,ta,c), sigma else - let c', sigma' = subst_closed_term_occ env sigma (AtOccs locc) a (EConstr.of_constr c) in + let c', sigma' = subst_closed_term_occ env sigma (AtOccs locc) a c in + let c' = EConstr.of_constr c' in mkLambda (na,ta,c'), sigma' let pattern_occs loccs_trm = { e_redfun = begin fun env sigma c -> + let open EConstr in let sigma = Sigma.to_evar_map sigma in - let abstr_trm, sigma = List.fold_right (abstract_scheme env) loccs_trm (EConstr.Unsafe.to_constr c,sigma) in + let abstr_trm, sigma = List.fold_right (abstract_scheme env sigma) loccs_trm (c,sigma) in try - let _ = Typing.unsafe_type_of env sigma (EConstr.of_constr abstr_trm) in - Sigma.Unsafe.of_pair (applist(abstr_trm, List.map snd loccs_trm), sigma) + let _ = Typing.unsafe_type_of env sigma abstr_trm in + Sigma.Unsafe.of_pair (EConstr.Unsafe.to_constr (applist(abstr_trm, List.map snd loccs_trm)), sigma) with Type_errors.TypeError (env',t) -> raise (ReductionTacticError (InvalidAbstraction (env,sigma,abstr_trm,(env',t)))) end } @@ -1190,28 +1218,31 @@ let check_not_primitive_record env ind = return name, B and t' *) let reduce_to_ind_gen allow_product env sigma t = + let open EConstr in let rec elimrec env t l = - let t = hnf_constr env sigma (EConstr.of_constr t) in - match kind_of_term (fst (decompose_app t)) with - | Ind ind-> (check_privacy env ind, it_mkProd_or_LetIn t l) + let t = hnf_constr env sigma t in + let t = EConstr.of_constr t in + match EConstr.kind sigma (EConstr.of_constr (fst (decompose_app_vect sigma t))) with + | Ind ind-> (check_privacy env ind, EConstr.Unsafe.to_constr (it_mkProd_or_LetIn t l)) | Prod (n,ty,t') -> let open Context.Rel.Declaration in if allow_product then - elimrec (push_rel (LocalAssum (n,ty)) env) t' ((LocalAssum (n,ty))::l) + elimrec (push_rel (local_assum (n,ty)) env) t' ((local_assum (n,ty))::l) else user_err (str"Not an inductive definition.") | _ -> (* Last chance: we allow to bypass the Opaque flag (as it was partially the case between V5.10 and V8.1 *) - let t' = whd_all env sigma (EConstr.of_constr t) in - match kind_of_term (fst (decompose_app t')) with - | Ind ind-> (check_privacy env ind, it_mkProd_or_LetIn t' l) + let t' = whd_all env sigma t in + let t' = EConstr.of_constr t' in + match EConstr.kind sigma (EConstr.of_constr (fst (decompose_app_vect sigma t'))) with + | Ind ind-> (check_privacy env ind, EConstr.Unsafe.to_constr (it_mkProd_or_LetIn t' l)) | _ -> user_err (str"Not an inductive product.") in elimrec env t [] -let reduce_to_quantified_ind x = reduce_to_ind_gen true x -let reduce_to_atomic_ind x = reduce_to_ind_gen false x +let reduce_to_quantified_ind env sigma c = reduce_to_ind_gen true env sigma (EConstr.of_constr c) +let reduce_to_atomic_ind env sigma c = reduce_to_ind_gen false env sigma (EConstr.of_constr c) let find_hnf_rectype env sigma t = let ind,t = reduce_to_atomic_ind env sigma t in @@ -1243,13 +1274,13 @@ let one_step_reduce env sigma c = | Reduced s' -> s' | NotReducible -> raise NotStepReducible with Redelimination -> raise NotStepReducible) - | _ when isEvalRef env (EConstr.Unsafe.to_constr x) -> - let ref,u = destEvalRefU (EConstr.Unsafe.to_constr x) in + | _ when isEvalRef env sigma x -> + let ref,u = destEvalRefU sigma x in (try fst (red_elim_const env sigma ref u stack) with Redelimination -> match reference_opt_value env sigma ref u with - | Some d -> (EConstr.of_constr d, stack) + | Some d -> (d, stack) | None -> raise NotStepReducible) | _ -> raise NotStepReducible @@ -1271,26 +1302,29 @@ let reduce_to_ref_gen allow_product env sigma ref t = else (* lazily reduces to match the head of [t] with the expected [ref] *) let rec elimrec env t l = - let c, _ = decompose_app_vect sigma (EConstr.of_constr t) in - match kind_of_term c with + let open EConstr in + let c, _ = decompose_app_vect sigma t in + let c = EConstr.of_constr c in + match EConstr.kind sigma c with | Prod (n,ty,t') -> if allow_product then let open Context.Rel.Declaration in - elimrec (push_rel (LocalAssum (n,t)) env) t' ((LocalAssum (n,ty))::l) + elimrec (push_rel (local_assum (n,t)) env) t' ((local_assum (n,ty))::l) else error_cannot_recognize ref | _ -> try - if eq_gr (global_of_constr c) ref + if eq_gr (global_of_constr (EConstr.to_constr sigma c)) ref then it_mkProd_or_LetIn t l else raise Not_found with Not_found -> try - let t' = nf_betaiota sigma (one_step_reduce env sigma (EConstr.of_constr t)) in + let t' = nf_betaiota sigma (one_step_reduce env sigma t) in + let t' = EConstr.of_constr t' in elimrec env t' l with NotStepReducible -> error_cannot_recognize ref in - elimrec env t [] + EConstr.Unsafe.to_constr (elimrec env t []) let reduce_to_quantified_ref = reduce_to_ref_gen true let reduce_to_atomic_ref = reduce_to_ref_gen false diff --git a/pretyping/tacred.mli b/pretyping/tacred.mli index f8dfe1adf..d32fcf491 100644 --- a/pretyping/tacred.mli +++ b/pretyping/tacred.mli @@ -17,7 +17,7 @@ open Locus open Univ type reduction_tactic_error = - InvalidAbstraction of env * evar_map * constr * (env * Type_errors.type_error) + InvalidAbstraction of env * evar_map * EConstr.constr * (env * Type_errors.type_error) exception ReductionTacticError of reduction_tactic_error @@ -58,10 +58,10 @@ val unfoldn : (occurrences * evaluable_global_reference) list -> reduction_function (** Fold *) -val fold_commands : constr list -> reduction_function +val fold_commands : EConstr.constr list -> reduction_function (** Pattern *) -val pattern_occs : (occurrences * constr) list -> e_reduction_function +val pattern_occs : (occurrences * EConstr.constr) list -> e_reduction_function (** Rem: Lazy strategies are defined in Reduction *) @@ -85,10 +85,10 @@ val reduce_to_quantified_ind : env -> evar_map -> types -> pinductive * types (** [reduce_to_quantified_ref env sigma ref t] try to put [t] in the form [t'=(x1:A1)..(xn:An)(ref args)] and fails with user error if not possible *) val reduce_to_quantified_ref : - env -> evar_map -> global_reference -> types -> types + env -> evar_map -> global_reference -> EConstr.types -> types val reduce_to_atomic_ref : - env -> evar_map -> global_reference -> types -> types + env -> evar_map -> global_reference -> EConstr.types -> types val find_hnf_rectype : env -> evar_map -> types -> pinductive * constr list diff --git a/pretyping/unification.ml b/pretyping/unification.ml index f418dc6a9..786cfd31f 100644 --- a/pretyping/unification.ml +++ b/pretyping/unification.ml @@ -1611,7 +1611,7 @@ let make_abstraction_core name (test,out) env sigma c ty occs check_occs concl = (push_named_context_val d sign,depdecls) | AllOccurrences, InHyp as occ -> let occ = if likefirst then LikeFirst else AtOccs occ in - let newdecl = replace_term_occ_decl_modulo occ test mkvarid d in + let newdecl = replace_term_occ_decl_modulo sigma occ test mkvarid d in if Context.Named.Declaration.equal d newdecl && not (indirectly_dependent sigma c d depdecls) then @@ -1622,7 +1622,7 @@ let make_abstraction_core name (test,out) env sigma c ty occs check_occs concl = (push_named_context_val newdecl sign, newdecl :: depdecls) | occ -> (* There are specific occurrences, hence not like first *) - let newdecl = replace_term_occ_decl_modulo (AtOccs occ) test mkvarid d in + let newdecl = replace_term_occ_decl_modulo sigma (AtOccs occ) test mkvarid d in (push_named_context_val newdecl sign, newdecl :: depdecls) in try let sign,depdecls = @@ -1632,7 +1632,7 @@ let make_abstraction_core name (test,out) env sigma c ty occs check_occs concl = | NoOccurrences -> concl | occ -> let occ = if likefirst then LikeFirst else AtOccs occ in - replace_term_occ_modulo occ test mkvarid (EConstr.of_constr concl) + replace_term_occ_modulo sigma occ test mkvarid (EConstr.of_constr concl) in let lastlhyp = if List.is_empty depdecls then None else Some (NamedDecl.get_id (List.last depdecls)) in diff --git a/proofs/redexpr.ml b/proofs/redexpr.ml index 19e72e697..d4a58da32 100644 --- a/proofs/redexpr.ml +++ b/proofs/redexpr.ml @@ -200,6 +200,9 @@ let out_arg = function let out_with_occurrences (occs,c) = (Locusops.occurrences_map (List.map out_arg) occs, c) +let out_with_occurrences' (occs,c) = + (Locusops.occurrences_map (List.map out_arg) occs, EConstr.of_constr c) + let e_red f = { e_redfun = fun env evm c -> Sigma.here (f env (Sigma.to_evar_map evm) c) evm } let head_style = false (* Turn to true to have a semantics where simpl @@ -239,8 +242,8 @@ let reduction_of_red_expr env = (e_red (strong_cbn (make_flag f)), DEFAULTcast) | Lazy f -> (e_red (clos_norm_flags (make_flag f)),DEFAULTcast) | Unfold ubinds -> (e_red (unfoldn (List.map out_with_occurrences ubinds)),DEFAULTcast) - | Fold cl -> (e_red (fold_commands cl),DEFAULTcast) - | Pattern lp -> (pattern_occs (List.map out_with_occurrences lp),DEFAULTcast) + | Fold cl -> (e_red (fold_commands (List.map EConstr.of_constr cl)),DEFAULTcast) + | Pattern lp -> (pattern_occs (List.map out_with_occurrences' lp),DEFAULTcast) | ExtraRedExpr s -> (try (e_red (String.Map.find s !reduction_tab),DEFAULTcast) with Not_found -> diff --git a/tactics/tactics.ml b/tactics/tactics.ml index 8fb47b994..e4503dab6 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -875,7 +875,7 @@ let normalise_vm_in_concl = reduct_in_concl (Redexpr.cbv_vm,VMcast) let unfold_in_concl loccname = reduct_in_concl (unfoldn loccname,REVERTcast) let unfold_in_hyp loccname = reduct_in_hyp (unfoldn loccname) let unfold_option loccname = reduct_option (unfoldn loccname,DEFAULTcast) -let pattern_option l = e_reduct_option (pattern_occs l,DEFAULTcast) +let pattern_option l = e_reduct_option (pattern_occs (List.map (on_snd EConstr.of_constr) l),DEFAULTcast) (* The main reduction function *) @@ -3165,7 +3165,7 @@ let atomize_param_of_ind_then (indref,nparams,_) hyp0 tac = let env = Proofview.Goal.env gl in let tmptyp0 = Tacmach.New.pf_get_hyp_typ hyp0 (Proofview.Goal.assume gl) in let reduce_to_quantified_ref = Tacmach.New.pf_apply reduce_to_quantified_ref gl in - let typ0 = reduce_to_quantified_ref indref tmptyp0 in + let typ0 = reduce_to_quantified_ref indref (EConstr.of_constr tmptyp0) in let prods, indtyp = decompose_prod_assum typ0 in let hd,argl = decompose_app indtyp in let env' = push_rel_context prods env in diff --git a/toplevel/himsg.ml b/toplevel/himsg.ml index b480fcd86..1d82d9ae1 100644 --- a/toplevel/himsg.ml +++ b/toplevel/himsg.ml @@ -1270,6 +1270,7 @@ let explain_pattern_matching_error env sigma = function let explain_reduction_tactic_error = function | Tacred.InvalidAbstraction (env,sigma,c,(env',e)) -> + let c = EConstr.to_constr sigma c in str "The abstracted term" ++ spc () ++ quote (pr_goal_concl_style_env env sigma c) ++ spc () ++ str "is not well typed." ++ fnl () ++ -- cgit v1.2.3 From 3b8acc174490878a3d0c9345e34a0ecb1d3abd66 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Mon, 7 Nov 2016 13:27:16 +0100 Subject: Typeclasses API using EConstr. --- engine/termops.ml | 9 ++++++++ engine/termops.mli | 2 ++ ltac/rewrite.ml | 6 +++--- pretyping/tacred.ml | 2 +- pretyping/typeclasses.ml | 44 ++++++++++++++++++++-------------------- pretyping/typeclasses.mli | 10 ++++----- pretyping/typeclasses_errors.ml | 1 + pretyping/typeclasses_errors.mli | 1 + tactics/class_tactics.ml | 14 ++++++------- tactics/tactics.ml | 1 + toplevel/classes.ml | 6 +++--- toplevel/himsg.ml | 5 +++-- toplevel/record.ml | 2 +- 13 files changed, 59 insertions(+), 44 deletions(-) diff --git a/engine/termops.ml b/engine/termops.ml index f191e2dc1..e5db3c085 100644 --- a/engine/termops.ml +++ b/engine/termops.ml @@ -818,6 +818,15 @@ let is_section_variable id = try let _ = Global.lookup_named id in true with Not_found -> false +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, Univ.Instance.empty + | _ -> raise Not_found + let isGlobalRef sigma c = match EConstr.kind sigma c with | Const _ | Ind _ | Construct _ | Var _ -> true diff --git a/engine/termops.mli b/engine/termops.mli index 4becca907..c90fdf9c2 100644 --- a/engine/termops.mli +++ b/engine/termops.mli @@ -264,6 +264,8 @@ val dependency_closure : env -> Evd.evar_map -> Context.Named.t -> Id.Set.t -> I (** Test if an identifier is the basename of a global reference *) val is_section_variable : Id.t -> bool +val global_of_constr : Evd.evar_map -> EConstr.constr -> Globnames.global_reference puniverses + val isGlobalRef : Evd.evar_map -> EConstr.t -> bool val is_template_polymorphic : env -> Evd.evar_map -> EConstr.t -> bool diff --git a/ltac/rewrite.ml b/ltac/rewrite.ml index 5703687c4..0091d0d0a 100644 --- a/ltac/rewrite.ml +++ b/ltac/rewrite.ml @@ -127,7 +127,7 @@ let app_poly_sort b = let find_class_proof proof_type proof_method env evars carrier relation = try let evars, goal = app_poly_check env evars proof_type [| carrier ; relation |] in - let evars', c = Typeclasses.resolve_one_typeclass env (goalevars evars) goal in + let evars', c = Typeclasses.resolve_one_typeclass env (goalevars evars) (EConstr.of_constr goal) in if extends_undefined (goalevars evars) evars' then raise Not_found else app_poly_check env (evars',cstrevars evars) proof_method [| carrier; relation; c |] with e when Logic.catchable_exception e -> raise Not_found @@ -367,7 +367,7 @@ end) = struct let evars, inst = app_poly env (evars,Evar.Set.empty) rewrite_relation_class [| evar; mkApp (c, params) |] in - let _ = Typeclasses.resolve_one_typeclass env' (goalevars evars) inst in + let _ = Typeclasses.resolve_one_typeclass env' (goalevars evars) (EConstr.of_constr inst) in Some (it_mkProd_or_LetIn t rels) with e when CErrors.noncritical e -> None) | _ -> None @@ -1937,7 +1937,7 @@ let default_morphism sign m = PropGlobal.build_signature (sigma, Evar.Set.empty) env t (fst sign) (snd sign) in let evars, morph = app_poly_check env evars PropGlobal.proper_type [| t; sign; m |] in - let evars, mor = resolve_one_typeclass env (goalevars evars) morph in + let evars, mor = resolve_one_typeclass env (goalevars evars) (EConstr.of_constr morph) in mor, proper_projection mor morph let add_setoid global binders a aeq t n = diff --git a/pretyping/tacred.ml b/pretyping/tacred.ml index 9997976c4..b729f3b9b 100644 --- a/pretyping/tacred.ml +++ b/pretyping/tacred.ml @@ -1314,7 +1314,7 @@ let reduce_to_ref_gen allow_product env sigma ref t = error_cannot_recognize ref | _ -> try - if eq_gr (global_of_constr (EConstr.to_constr sigma c)) ref + if eq_gr (fst (global_of_constr sigma c)) ref then it_mkProd_or_LetIn t l else raise Not_found with Not_found -> diff --git a/pretyping/typeclasses.ml b/pretyping/typeclasses.ml index 11f71ee02..a970c434f 100644 --- a/pretyping/typeclasses.ml +++ b/pretyping/typeclasses.ml @@ -136,23 +136,24 @@ let typeclass_univ_instance (cl,u') = let class_info c = try Refmap.find c !classes - with Not_found -> not_a_class (Global.env()) (printable_constr_of_global c) + with Not_found -> not_a_class (Global.env()) (EConstr.of_constr (printable_constr_of_global c)) -let global_class_of_constr env c = - try let gr, u = Universes.global_of_constr c in +let global_class_of_constr env sigma c = + try let gr, u = Termops.global_of_constr sigma c in class_info gr, u with Not_found -> not_a_class env c -let dest_class_app env c = - let cl, args = decompose_app c in - global_class_of_constr env cl, args +let dest_class_app env sigma c = + let cl, args = EConstr.decompose_app sigma c in + global_class_of_constr env sigma cl, (List.map EConstr.Unsafe.to_constr args) -let dest_class_arity env c = - let rels, c = decompose_prod_assum c in - rels, dest_class_app env c +let dest_class_arity env sigma c = + let open EConstr in + let rels, c = decompose_prod_assum sigma c in + rels, dest_class_app env sigma c -let class_of_constr c = - try Some (dest_class_arity (Global.env ()) c) +let class_of_constr sigma c = + try Some (dest_class_arity (Global.env ()) sigma c) with e when CErrors.noncritical e -> None let is_class_constr c = @@ -161,15 +162,14 @@ let is_class_constr c = with Not_found -> false let rec is_class_type evd c = - let c, args = decompose_app c in - match kind_of_term c with + let c, _ = Termops.decompose_app_vect evd c in + match EConstr.kind evd (EConstr.of_constr c) with | Prod (_, _, t) -> is_class_type evd t - | Evar (e, _) when Evd.is_defined evd e -> - is_class_type evd (EConstr.Unsafe.to_constr (Evarutil.whd_head_evar evd (EConstr.of_constr c))) + | Cast (t, _, _) -> is_class_type evd t | _ -> is_class_constr c let is_class_evar evd evi = - is_class_type evd evi.Evd.evar_concl + is_class_type evd (EConstr.of_constr evi.Evd.evar_concl) (* * classes persistent object @@ -222,7 +222,7 @@ let discharge_class (_,cl) = let discharge_context ctx' subst (grs, ctx) = let grs' = let newgrs = List.map (fun decl -> - match decl |> RelDecl.get_type |> class_of_constr with + match decl |> RelDecl.get_type |> EConstr.of_constr |> class_of_constr Evd.empty with | None -> None | Some (_, ((tc,_), _)) -> Some (tc.cl_impl, true)) ctx' @@ -270,7 +270,7 @@ let add_class cl = let check_instance env sigma c = try let (evd, c) = resolve_one_typeclass env sigma - (Retyping.get_type_of env sigma c) in + (EConstr.of_constr (Retyping.get_type_of env sigma c)) in not (Evd.has_undefined evd) with e when CErrors.noncritical e -> false @@ -282,10 +282,10 @@ let build_subclasses ~check env sigma glob pri = Nameops.add_suffix _id ("_subinstance_" ^ string_of_int !i)) in let ty, ctx = Global.type_of_global_in_context env glob in + let ty = EConstr.of_constr ty in let sigma = Evd.merge_context_set Evd.univ_rigid sigma (Univ.ContextSet.of_context ctx) in let rec aux pri c ty path = - let ty = Evarutil.nf_evar sigma ty in - match class_of_constr ty with + match class_of_constr sigma ty with | None -> [] | Some (rels, ((tc,u), args)) -> let instapp = @@ -313,7 +313,7 @@ let build_subclasses ~check env sigma glob pri = let declare_proj hints (cref, pri, body) = let path' = cref :: path in let ty = Retyping.get_type_of env sigma (EConstr.of_constr body) in - let rest = aux pri body ty path' in + let rest = aux pri body (EConstr.of_constr ty) path' in hints @ (path', pri, body) :: rest in List.fold_left declare_proj [] projs in @@ -406,7 +406,7 @@ let remove_instance i = let declare_instance pri local glob = let ty = Global.type_of_global_unsafe glob in - match class_of_constr ty with + match class_of_constr Evd.empty (EConstr.of_constr ty) with | Some (rels, ((tc,_), args) as _cl) -> add_instance (new_instance tc pri (not local) (Flags.use_polymorphic_flag ()) glob) (* let path, hints = build_subclasses (not local) (Global.env ()) Evd.empty glob in *) diff --git a/pretyping/typeclasses.mli b/pretyping/typeclasses.mli index 2530f5dfa..ec36c57e0 100644 --- a/pretyping/typeclasses.mli +++ b/pretyping/typeclasses.mli @@ -61,13 +61,13 @@ val class_info : global_reference -> typeclass (** raises a UserError if not a c (** These raise a UserError if not a class. Caution: the typeclass structures is not instantiated w.r.t. the universe instance. This is done separately by typeclass_univ_instance. *) -val dest_class_app : env -> constr -> typeclass puniverses * constr list +val dest_class_app : env -> evar_map -> EConstr.constr -> typeclass puniverses * constr list (** Get the instantiated typeclass structure for a given universe instance. *) val typeclass_univ_instance : typeclass puniverses -> typeclass puniverses (** Just return None if not a class *) -val class_of_constr : constr -> (Context.Rel.t * (typeclass puniverses * constr list)) option +val class_of_constr : evar_map -> EConstr.constr -> (Context.Rel.t * (typeclass puniverses * constr list)) option val instance_impl : instance -> global_reference @@ -99,11 +99,11 @@ val mark_unresolvables : ?filter:evar_filter -> evar_map -> evar_map val mark_resolvables : ?filter:evar_filter -> evar_map -> evar_map val mark_resolvable : evar_info -> evar_info val is_class_evar : evar_map -> evar_info -> bool -val is_class_type : evar_map -> types -> bool +val is_class_type : evar_map -> EConstr.types -> bool val resolve_typeclasses : ?fast_path:bool -> ?filter:evar_filter -> ?unique:bool -> ?split:bool -> ?fail:bool -> env -> evar_map -> evar_map -val resolve_one_typeclass : ?unique:bool -> env -> evar_map -> types -> open_constr +val resolve_one_typeclass : ?unique:bool -> env -> evar_map -> EConstr.types -> open_constr val set_typeclass_transparency_hook : (evaluable_global_reference -> bool (*local?*) -> bool -> unit) Hook.t val set_typeclass_transparency : evaluable_global_reference -> bool -> bool -> unit @@ -120,7 +120,7 @@ val add_instance_hint : global_reference_or_constr -> global_reference list -> val remove_instance_hint : global_reference -> unit val solve_all_instances_hook : (env -> evar_map -> evar_filter -> bool -> bool -> bool -> evar_map) Hook.t -val solve_one_instance_hook : (env -> evar_map -> types -> bool -> open_constr) Hook.t +val solve_one_instance_hook : (env -> evar_map -> EConstr.types -> bool -> open_constr) Hook.t val declare_instance : int option -> bool -> global_reference -> unit diff --git a/pretyping/typeclasses_errors.ml b/pretyping/typeclasses_errors.ml index b1dfb19a0..2db0e9e88 100644 --- a/pretyping/typeclasses_errors.ml +++ b/pretyping/typeclasses_errors.ml @@ -9,6 +9,7 @@ (*i*) open Names open Term +open EConstr open Environ open Constrexpr open Globnames diff --git a/pretyping/typeclasses_errors.mli b/pretyping/typeclasses_errors.mli index ee76f6383..9bd430e4d 100644 --- a/pretyping/typeclasses_errors.mli +++ b/pretyping/typeclasses_errors.mli @@ -9,6 +9,7 @@ open Loc open Names open Term +open EConstr open Environ open Constrexpr open Globnames diff --git a/tactics/class_tactics.ml b/tactics/class_tactics.ml index bef43d20b..ff7dbfa91 100644 --- a/tactics/class_tactics.ml +++ b/tactics/class_tactics.ml @@ -478,9 +478,9 @@ let is_Prop env sigma concl = | Sort (Prop Null) -> true | _ -> false -let is_unique env concl = +let is_unique env sigma concl = try - let (cl,u), args = dest_class_app env concl in + let (cl,u), args = dest_class_app env sigma concl in cl.cl_unique with e when CErrors.noncritical e -> false @@ -675,7 +675,7 @@ module V85 = struct let tacgl = {it = gl; sigma = s;} in let secvars = secvars_of_hyps (Environ.named_context_of_val (Goal.V82.hyps s gl)) in let poss = e_possible_resolve hints info.hints secvars info.only_classes s concl in - let unique = is_unique env concl in + let unique = is_unique env s (EConstr.of_constr concl) in let rec aux i foundone = function | (tac, _, extern, name, pp) :: tl -> let derivs = path_derivate info.auto_cut name in @@ -997,7 +997,7 @@ module Search = struct let concl = Goal.concl gl in let sigma = Goal.sigma gl in let s = Sigma.to_evar_map sigma in - let unique = not info.search_dep || is_unique env concl in + let unique = not info.search_dep || is_unique env s (EConstr.of_constr concl) in let backtrack = needs_backtrack env s unique concl in if !typeclasses_debug > 0 then Feedback.msg_debug @@ -1071,7 +1071,7 @@ module Search = struct try let evi = Evd.find_undefined sigma ev in if info.search_only_classes then - Some (ev, is_class_type sigma (Evd.evar_concl evi)) + Some (ev, is_class_evar sigma evi) else Some (ev, true) with Not_found -> None in @@ -1351,7 +1351,7 @@ let error_unresolvable env comp evd = | Some s -> Evar.Set.mem ev s in let fold ev evi (found, accu) = - let ev_class = class_of_constr evi.evar_concl in + let ev_class = class_of_constr evd (EConstr.of_constr evi.evar_concl) in if not (Option.is_empty ev_class) && is_part ev then (* focus on one instance if only one was searched for *) if not found then (true, Some ev) @@ -1481,7 +1481,7 @@ let resolve_one_typeclass env ?(sigma=Evd.empty) gl unique = let _ = Hook.set Typeclasses.solve_one_instance_hook - (fun x y z w -> resolve_one_typeclass x ~sigma:y z w) + (fun x y z w -> resolve_one_typeclass x ~sigma:y (EConstr.Unsafe.to_constr z) w) (** Take the head of the arity of a constr. Used in the partial application tactic. *) diff --git a/tactics/tactics.ml b/tactics/tactics.ml index e4503dab6..a6bc805bd 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -1672,6 +1672,7 @@ let solve_remaining_apply_goals = let env = Proofview.Goal.env gl in let evd = Sigma.to_evar_map sigma in let concl = Proofview.Goal.concl gl in + let concl = EConstr.of_constr concl in if Typeclasses.is_class_type evd concl then let evd', c' = Typeclasses.resolve_one_typeclass env evd concl in let tac = Refine.refine ~unsafe:true { run = fun h -> Sigma.here c' h } in diff --git a/toplevel/classes.ml b/toplevel/classes.ml index ad4a13c21..26e29540c 100644 --- a/toplevel/classes.ml +++ b/toplevel/classes.ml @@ -65,7 +65,7 @@ let existing_instance glob g pri = let c = global g in let instance = Global.type_of_global_unsafe c in let _, r = decompose_prod_assum instance in - match class_of_constr r with + match class_of_constr Evd.empty (EConstr.of_constr r) with | Some (_, ((tc,u), _)) -> add_instance (new_instance tc pri glob (*FIXME*) (Flags.use_polymorphic_flag ()) c) | None -> user_err ~loc:(loc_of_reference g) @@ -155,7 +155,7 @@ let new_instance ?(abstract=false) ?(global=false) ?(refine= !refine_instance) p let imps = imps @ Impargs.lift_implicits len imps' in let ctx', c = decompose_prod_assum c' in let ctx'' = ctx' @ ctx in - let k, args = Typeclasses.dest_class_app (push_rel_context ctx'' env) c in + let k, args = Typeclasses.dest_class_app (push_rel_context ctx'' env) !evars (EConstr.of_constr c) in let cl, u = Typeclasses.typeclass_univ_instance k in let _, args = List.fold_right (fun decl (args, args') -> @@ -378,7 +378,7 @@ let context poly l = let () = uctx := Univ.ContextSet.empty in let decl = (ParameterEntry (None,poly,(t,ctx),None), IsAssumption Logical) in let cst = Declare.declare_constant ~internal:Declare.InternalTacticRequest id decl in - match class_of_constr t with + match class_of_constr !evars (EConstr.of_constr t) with | Some (rels, ((tc,_), args) as _cl) -> add_instance (Typeclasses.new_instance tc None false (*FIXME*) poly (ConstRef cst)); diff --git a/toplevel/himsg.ml b/toplevel/himsg.ml index 1d82d9ae1..df1f47e33 100644 --- a/toplevel/himsg.ml +++ b/toplevel/himsg.ml @@ -584,7 +584,7 @@ let rec explain_evar_kind env sigma evk ty = function (pr_lconstr_env env sigma ty') (snd evi.evar_source) let explain_typeclass_resolution env sigma evi k = - match Typeclasses.class_of_constr evi.evar_concl with + match Typeclasses.class_of_constr sigma (EConstr.of_constr evi.evar_concl) with | Some _ -> let env = Evd.evar_filtered_env evi in fnl () ++ str "Could not find an instance for " ++ @@ -597,7 +597,7 @@ let explain_placeholder_kind env sigma c e = | Some (SeveralInstancesFound n) -> strbrk " (several distinct possible type class instances found)" | None -> - match Typeclasses.class_of_constr c with + match Typeclasses.class_of_constr sigma (EConstr.of_constr c) with | Some _ -> strbrk " (no type class instance found)" | _ -> mt () @@ -1012,6 +1012,7 @@ let explain_module_internalization_error = function (* Typeclass errors *) let explain_not_a_class env c = + let c = EConstr.to_constr Evd.empty c in pr_constr_env env Evd.empty c ++ str" is not a declared type class." let explain_unbound_method env cid id = diff --git a/toplevel/record.ml b/toplevel/record.ml index 5f2b99f90..dc49c5385 100644 --- a/toplevel/record.ml +++ b/toplevel/record.ml @@ -484,7 +484,7 @@ let declare_class finite def poly ctx id idbuild paramimpls params arity in let ctx_context = List.map (fun decl -> - match Typeclasses.class_of_constr (RelDecl.get_type decl) with + match Typeclasses.class_of_constr Evd.empty (EConstr.of_constr (RelDecl.get_type decl)) with | Some (_, ((cl,_), _)) -> Some (cl.cl_impl, true) | None -> None) params, params -- cgit v1.2.3 From ce2b509734f3b70494a0a35b0b4eda593c1c8eb6 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Mon, 7 Nov 2016 19:07:16 +0100 Subject: Classops API using EConstr. --- interp/notation.ml | 2 +- pretyping/classops.ml | 16 ++++++++-------- pretyping/classops.mli | 12 ++++++------ pretyping/coercion.ml | 8 ++++---- toplevel/class.ml | 8 ++++---- 5 files changed, 23 insertions(+), 23 deletions(-) diff --git a/interp/notation.ml b/interp/notation.ml index 88ae4695b..d264a1904 100644 --- a/interp/notation.ml +++ b/interp/notation.ml @@ -585,7 +585,7 @@ let scope_class_compare : scope_class -> scope_class -> int = cl_typ_ord let compute_scope_class t = - let (cl,_,_) = find_class_type Evd.empty t in + let (cl,_,_) = find_class_type Evd.empty (EConstr.of_constr t) in cl module ScopeClassOrd = diff --git a/pretyping/classops.ml b/pretyping/classops.ml index 577f41a7d..753127357 100644 --- a/pretyping/classops.ml +++ b/pretyping/classops.ml @@ -193,7 +193,7 @@ let coercion_exists coe = CoeTypMap.mem coe !coercion_tab let find_class_type sigma t = let inj = EConstr.Unsafe.to_constr in - let t', args = Reductionops.whd_betaiotazeta_stack sigma (EConstr.of_constr t) in + let t', args = Reductionops.whd_betaiotazeta_stack sigma t in match EConstr.kind sigma t' with | Var id -> CL_SECVAR id, Univ.Instance.empty, List.map inj args | Const (sp,u) -> CL_CONST sp, u, List.map inj args @@ -215,7 +215,7 @@ let subst_cl_typ subst ct = match ct with | CL_CONST c -> let c',t = subst_con_kn subst c in if c' == c then ct else - pi1 (find_class_type Evd.empty t) + pi1 (find_class_type Evd.empty (EConstr.of_constr t)) | CL_IND i -> let i' = subst_ind subst i in if i' == i then ct else CL_IND i' @@ -231,10 +231,10 @@ let class_of env sigma t = try let (cl, u, args) = find_class_type sigma t in let (i, { cl_param = n1 } ) = class_info cl in - (t, n1, i, u, args) + (EConstr.Unsafe.to_constr t, n1, i, u, args) with Not_found -> - let t = Tacred.hnf_constr env sigma (EConstr.of_constr t) in - let (cl, u, args) = find_class_type sigma t in + let t = Tacred.hnf_constr env sigma t in + let (cl, u, args) = find_class_type sigma (EConstr.of_constr t) in let (i, { cl_param = n1 } ) = class_info cl in (t, n1, i, u, args) in @@ -274,11 +274,11 @@ let apply_on_class_of env sigma t cont = let (cl,u,args) = find_class_type sigma t in let (i, { cl_param = n1 } ) = class_info cl in if not (Int.equal (List.length args) n1) then raise Not_found; - t, cont i + EConstr.Unsafe.to_constr t, cont i with Not_found -> (* Is it worth to be more incremental on the delta steps? *) - let t = Tacred.hnf_constr env sigma (EConstr.of_constr t) in - let (cl, u, args) = find_class_type sigma t in + let t = Tacred.hnf_constr env sigma t in + let (cl, u, args) = find_class_type sigma (EConstr.of_constr t) in let (i, { cl_param = n1 } ) = class_info cl in if not (Int.equal (List.length args) n1) then raise Not_found; t, cont i diff --git a/pretyping/classops.mli b/pretyping/classops.mli index d509739cf..4b8a2c1c0 100644 --- a/pretyping/classops.mli +++ b/pretyping/classops.mli @@ -59,15 +59,15 @@ val class_info_from_index : cl_index -> cl_typ * cl_info_typ (** [find_class_type env sigma c] returns the head reference of [c], its universe instance and its arguments *) -val find_class_type : evar_map -> types -> cl_typ * Univ.universe_instance * constr list +val find_class_type : evar_map -> EConstr.types -> cl_typ * Univ.universe_instance * constr list (** raises [Not_found] if not convertible to a class *) -val class_of : env -> evar_map -> types -> types * cl_index +val class_of : env -> evar_map -> EConstr.types -> types * cl_index (** raises [Not_found] if not mapped to a class *) val inductive_class_of : inductive -> cl_index -val class_args_of : env -> evar_map -> types -> constr list +val class_args_of : env -> evar_map -> EConstr.types -> constr list (** {6 [declare_coercion] adds a coercion in the graph of coercion paths } *) val declare_coercion : @@ -84,11 +84,11 @@ val coercion_value : coe_index -> (unsafe_judgment * bool * bool) Univ.in_univer (** @raise Not_found in the following functions when no path exists *) val lookup_path_between_class : cl_index * cl_index -> inheritance_path -val lookup_path_between : env -> evar_map -> types * types -> +val lookup_path_between : env -> evar_map -> EConstr.types * EConstr.types -> types * types * inheritance_path -val lookup_path_to_fun_from : env -> evar_map -> types -> +val lookup_path_to_fun_from : env -> evar_map -> EConstr.types -> types * inheritance_path -val lookup_path_to_sort_from : env -> evar_map -> types -> +val lookup_path_to_sort_from : env -> evar_map -> EConstr.types -> types * inheritance_path val lookup_pattern_path_between : env -> inductive * inductive -> (constructor * int) list diff --git a/pretyping/coercion.ml b/pretyping/coercion.ml index 04e235cc5..90cd3b60b 100644 --- a/pretyping/coercion.ml +++ b/pretyping/coercion.ml @@ -358,7 +358,7 @@ let apply_coercion env sigma p hj typ_cl = (fun (ja,typ_cl,sigma) i -> let ((fv,isid,isproj),ctx) = coercion_value i in let sigma = Evd.merge_context_set Evd.univ_flexible sigma ctx in - let argl = (class_args_of env sigma typ_cl)@[ja.uj_val] in + let argl = (class_args_of env sigma (EConstr.of_constr typ_cl))@[ja.uj_val] in let sigma, jres = apply_coercion_args env sigma true isproj argl fv in @@ -382,7 +382,7 @@ let inh_app_fun_core env evd j = (evd',{ uj_val = j.uj_val; uj_type = EConstr.Unsafe.to_constr t }) | _ -> try let t,p = - lookup_path_to_fun_from env evd j.uj_type in + lookup_path_to_fun_from env evd (EConstr.of_constr j.uj_type) in apply_coercion env evd p j t with Not_found | NoCoercion -> if Flags.is_program_mode () then @@ -407,7 +407,7 @@ let inh_app_fun resolve_tc env evd j = let inh_tosort_force loc env evd j = try - let t,p = lookup_path_to_sort_from env evd j.uj_type in + let t,p = lookup_path_to_sort_from env evd (EConstr.of_constr j.uj_type) in let evd,j1 = apply_coercion env evd p j t in let j2 = on_judgment_type (whd_evar evd) j1 in (evd,type_judgment env j2) @@ -448,7 +448,7 @@ let inh_coerce_to_fail env evd rigidonly v t c1 = else let evd, v', t' = try - let t2,t1,p = lookup_path_between env evd (t,c1) in + let t2,t1,p = lookup_path_between env evd (EConstr.of_constr t,EConstr.of_constr c1) in match v with | Some v -> let evd,j = diff --git a/toplevel/class.ml b/toplevel/class.ml index 86fd4d9a2..6295eb336 100644 --- a/toplevel/class.ml +++ b/toplevel/class.ml @@ -120,7 +120,7 @@ let get_source lp source = let (cl1,u1,lv1) = match lp with | [] -> raise Not_found - | t1::_ -> find_class_type Evd.empty t1 + | t1::_ -> find_class_type Evd.empty (EConstr.of_constr t1) in (cl1,u1,lv1,1) | Some cl -> @@ -128,7 +128,7 @@ let get_source lp source = | [] -> raise Not_found | t1::lt -> try - let cl1,u1,lv1 = find_class_type Evd.empty t1 in + let cl1,u1,lv1 = find_class_type Evd.empty (EConstr.of_constr t1) in if cl_typ_eq cl cl1 then cl1,u1,lv1,(List.length lt+1) else raise Not_found with Not_found -> aux lt @@ -138,7 +138,7 @@ let get_target t ind = if (ind > 1) then CL_FUN else - match pi1 (find_class_type Evd.empty t) with + match pi1 (find_class_type Evd.empty (EConstr.of_constr t)) with | CL_CONST p when Environ.is_projection p (Global.env ()) -> CL_PROJ p | x -> x @@ -215,7 +215,7 @@ let build_id_coercion idf_opt source poly = match idf_opt with | Some idf -> idf | None -> - let cl,u,_ = find_class_type sigma t in + let cl,u,_ = find_class_type sigma (EConstr.of_constr t) in Id.of_string ("Id_"^(ident_key_of_class source)^"_"^ (ident_key_of_class cl)) in -- cgit v1.2.3 From e4f066238799a4598817dfeab8a044760ab670de Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Mon, 7 Nov 2016 20:33:06 +0100 Subject: Coercion API using EConstr. --- pretyping/cases.ml | 5 +- pretyping/classops.ml | 19 +++--- pretyping/classops.mli | 15 ++-- pretyping/coercion.ml | 174 +++++++++++++++++++++++++---------------------- pretyping/coercion.mli | 8 +-- pretyping/pretyping.ml | 6 +- pretyping/program.ml | 3 +- pretyping/program.mli | 2 +- pretyping/unification.ml | 2 +- proofs/clenv.ml | 2 +- proofs/refine.ml | 2 +- toplevel/class.ml | 2 +- 12 files changed, 128 insertions(+), 112 deletions(-) diff --git a/pretyping/cases.ml b/pretyping/cases.ml index 882c052f6..96c61647c 100644 --- a/pretyping/cases.ml +++ b/pretyping/cases.ml @@ -396,7 +396,7 @@ let adjust_tomatch_to_pattern pb ((current,typ),deps,dep) = current else (evd_comb2 (Coercion.inh_conv_coerce_to true Loc.ghost pb.env) - pb.evdref (make_judge current typ) indt).uj_val in + pb.evdref (make_judge current typ) (EConstr.of_constr indt)).uj_val in let sigma = !(pb.evdref) in (current,try_find_ind pb.env sigma indt names)) | _ -> (current,tmtyp) @@ -1867,7 +1867,7 @@ let extract_arity_signature ?(dolift=true) env0 tomatchl tmsign = let inh_conv_coerce_to_tycon loc env evdref j tycon = match tycon with | Some p -> - let (evd',j) = Coercion.inh_conv_coerce_to true loc env !evdref j p in + let (evd',j) = Coercion.inh_conv_coerce_to true loc env !evdref j (EConstr.of_constr p) in evdref := evd'; j | None -> j @@ -2013,6 +2013,7 @@ let eq_id avoid id = let hid' = next_ident_away hid avoid in hid' +let papp evdref gr args = EConstr.Unsafe.to_constr (papp evdref gr (Array.map EConstr.of_constr args)) let mk_eq evdref typ x y = papp evdref coq_eq_ind [| typ; x ; y |] let mk_eq_refl evdref typ x = papp evdref coq_eq_refl [| typ; x |] let mk_JMeq evdref typ x typ' y = diff --git a/pretyping/classops.ml b/pretyping/classops.ml index 753127357..ad43bf322 100644 --- a/pretyping/classops.ml +++ b/pretyping/classops.ml @@ -192,14 +192,13 @@ let coercion_exists coe = CoeTypMap.mem coe !coercion_tab (* find_class_type : evar_map -> constr -> cl_typ * universe_list * constr list *) let find_class_type sigma t = - let inj = EConstr.Unsafe.to_constr in let t', args = Reductionops.whd_betaiotazeta_stack sigma t in match EConstr.kind sigma t' with - | Var id -> CL_SECVAR id, Univ.Instance.empty, List.map inj args - | Const (sp,u) -> CL_CONST sp, u, List.map inj args + | Var id -> CL_SECVAR id, Univ.Instance.empty, args + | Const (sp,u) -> CL_CONST sp, u, args | Proj (p, c) when not (Projection.unfolded p) -> - CL_PROJ (Projection.constant p), Univ.Instance.empty, List.map inj (c :: args) - | Ind (ind_sp,u) -> CL_IND ind_sp, u, List.map inj args + CL_PROJ (Projection.constant p), Univ.Instance.empty, (c :: args) + | Ind (ind_sp,u) -> CL_IND ind_sp, u, args | Prod (_,_,_) -> CL_FUN, Univ.Instance.empty, [] | Sort _ -> CL_SORT, Univ.Instance.empty, [] | _ -> raise Not_found @@ -231,10 +230,11 @@ let class_of env sigma t = try let (cl, u, args) = find_class_type sigma t in let (i, { cl_param = n1 } ) = class_info cl in - (EConstr.Unsafe.to_constr t, n1, i, u, args) + (t, n1, i, u, args) with Not_found -> let t = Tacred.hnf_constr env sigma t in - let (cl, u, args) = find_class_type sigma (EConstr.of_constr t) in + let t = EConstr.of_constr t in + let (cl, u, args) = find_class_type sigma t in let (i, { cl_param = n1 } ) = class_info cl in (t, n1, i, u, args) in @@ -274,11 +274,12 @@ let apply_on_class_of env sigma t cont = let (cl,u,args) = find_class_type sigma t in let (i, { cl_param = n1 } ) = class_info cl in if not (Int.equal (List.length args) n1) then raise Not_found; - EConstr.Unsafe.to_constr t, cont i + t, cont i with Not_found -> (* Is it worth to be more incremental on the delta steps? *) let t = Tacred.hnf_constr env sigma t in - let (cl, u, args) = find_class_type sigma (EConstr.of_constr t) in + let t = EConstr.of_constr t in + let (cl, u, args) = find_class_type sigma t in let (i, { cl_param = n1 } ) = class_info cl in if not (Int.equal (List.length args) n1) then raise Not_found; t, cont i diff --git a/pretyping/classops.mli b/pretyping/classops.mli index 4b8a2c1c0..9fb70534f 100644 --- a/pretyping/classops.mli +++ b/pretyping/classops.mli @@ -8,6 +8,7 @@ open Names open Term +open EConstr open Evd open Environ open Mod_subst @@ -59,15 +60,15 @@ val class_info_from_index : cl_index -> cl_typ * cl_info_typ (** [find_class_type env sigma c] returns the head reference of [c], its universe instance and its arguments *) -val find_class_type : evar_map -> EConstr.types -> cl_typ * Univ.universe_instance * constr list +val find_class_type : evar_map -> types -> cl_typ * Univ.universe_instance * constr list (** raises [Not_found] if not convertible to a class *) -val class_of : env -> evar_map -> EConstr.types -> types * cl_index +val class_of : env -> evar_map -> types -> types * cl_index (** raises [Not_found] if not mapped to a class *) val inductive_class_of : inductive -> cl_index -val class_args_of : env -> evar_map -> EConstr.types -> constr list +val class_args_of : env -> evar_map -> types -> constr list (** {6 [declare_coercion] adds a coercion in the graph of coercion paths } *) val declare_coercion : @@ -84,11 +85,11 @@ val coercion_value : coe_index -> (unsafe_judgment * bool * bool) Univ.in_univer (** @raise Not_found in the following functions when no path exists *) val lookup_path_between_class : cl_index * cl_index -> inheritance_path -val lookup_path_between : env -> evar_map -> EConstr.types * EConstr.types -> +val lookup_path_between : env -> evar_map -> types * types -> types * types * inheritance_path -val lookup_path_to_fun_from : env -> evar_map -> EConstr.types -> +val lookup_path_to_fun_from : env -> evar_map -> types -> types * inheritance_path -val lookup_path_to_sort_from : env -> evar_map -> EConstr.types -> +val lookup_path_to_sort_from : env -> evar_map -> types -> types * inheritance_path val lookup_pattern_path_between : env -> inductive * inductive -> (constructor * int) list @@ -104,7 +105,7 @@ val install_path_printer : val string_of_class : cl_typ -> string val pr_class : cl_typ -> std_ppcmds val pr_cl_index : cl_index -> std_ppcmds -val get_coercion_value : coe_index -> constr +val get_coercion_value : coe_index -> Constr.t val inheritance_graph : unit -> ((cl_index * cl_index) * inheritance_path) list val classes : unit -> cl_typ list val coercions : unit -> coe_index list diff --git a/pretyping/coercion.ml b/pretyping/coercion.ml index 90cd3b60b..cc121a96d 100644 --- a/pretyping/coercion.ml +++ b/pretyping/coercion.ml @@ -48,29 +48,30 @@ exception NoCoercionNoUnifier of evar_map * unification_error (* Here, funj is a coercion therefore already typed in global context *) let apply_coercion_args env evd check isproj argl funj = + let open EConstr in let evdref = ref evd in let rec apply_rec acc typ = function | [] -> if isproj then - let cst = fst (destConst (j_val funj)) in + let cst = fst (destConst !evdref (EConstr.of_constr (j_val funj))) in let p = Projection.make cst false in let pb = lookup_projection p env in let args = List.skipn pb.Declarations.proj_npars argl in let hd, tl = match args with hd :: tl -> hd, tl | [] -> assert false in - { uj_val = applist (mkProj (p, hd), tl); - uj_type = typ } + { uj_val = EConstr.Unsafe.to_constr (applist (mkProj (p, hd), tl)); + uj_type = EConstr.Unsafe.to_constr typ } else - { uj_val = applist (j_val funj,argl); - uj_type = typ } + { uj_val = EConstr.Unsafe.to_constr (applist (EConstr.of_constr (j_val funj),argl)); + uj_type = EConstr.Unsafe.to_constr typ } | h::restl -> (* On devrait pouvoir s'arranger pour qu'on n'ait pas a faire hnf_constr *) - match kind_of_term (whd_all env evd (EConstr.of_constr typ)) with + match EConstr.kind !evdref (EConstr.of_constr (whd_all env !evdref typ)) with | Prod (_,c1,c2) -> - if check && not (e_cumul env evdref (EConstr.of_constr (Retyping.get_type_of env evd (EConstr.of_constr h))) (EConstr.of_constr c1)) then + if check && not (e_cumul env evdref (EConstr.of_constr (Retyping.get_type_of env !evdref h)) c1) then raise NoCoercion; - apply_rec (h::acc) (subst1 h c2) restl + apply_rec (h::acc) (Vars.subst1 h c2) restl | _ -> anomaly (Pp.str "apply_coercion_args") in - let res = apply_rec [] funj.uj_type argl in + let res = apply_rec [] (EConstr.of_constr funj.uj_type) argl in !evdref, res (* appliquer le chemin de coercions de patterns p *) @@ -92,17 +93,17 @@ open Program let make_existential loc ?(opaque = not (get_proofs_transparency ())) env evdref c = let src = (loc, Evar_kinds.QuestionMark (Evar_kinds.Define opaque)) in - Evarutil.e_new_evar env evdref ~src c + EConstr.of_constr (Evarutil.e_new_evar env evdref ~src (EConstr.Unsafe.to_constr c)) let app_opt env evdref f t = - whd_betaiota !evdref (EConstr.of_constr (app_opt f t)) + EConstr.of_constr (whd_betaiota !evdref (app_opt f t)) let pair_of_array a = (a.(0), a.(1)) -let disc_subset x = - match kind_of_term x with +let disc_subset sigma x = + match EConstr.kind sigma x with | App (c, l) -> - (match kind_of_term c with + (match EConstr.kind sigma c with Ind (i,_) -> let len = Array.length l in let sigty = delayed_force sig_typ in @@ -120,19 +121,25 @@ let hnf env evd c = whd_all env evd c let hnf_nodelta env evd c = whd_betaiota evd c let lift_args n sign = + let open EConstr in let rec liftrec k = function - | t::sign -> liftn n k t :: (liftrec (k-1) sign) + | t::sign -> Vars.liftn n k t :: (liftrec (k-1) sign) | [] -> [] in liftrec (List.length sign) sign +let local_assum (na, t) = + let open Context.Rel.Declaration in + LocalAssum (na, EConstr.Unsafe.to_constr t) + let mu env evdref t = let rec aux v = - let v' = hnf env !evdref (EConstr.of_constr v) in - match disc_subset v' with + let v' = hnf env !evdref v in + let v' = EConstr.of_constr v' in + match disc_subset !evdref v' with | Some (u, p) -> let f, ct = aux u in - let p = hnf_nodelta env !evdref (EConstr.of_constr p) in + let p = EConstr.of_constr (hnf_nodelta env !evdref p) in (Some (fun x -> app_opt env evdref f (papp evdref sig_proj1 [| u; p; x |])), @@ -140,21 +147,25 @@ let mu env evdref t = | None -> (None, v) in aux t -and coerce loc env evdref (x : Term.constr) (y : Term.constr) - : (Term.constr -> Term.constr) option +and coerce loc env evdref (x : EConstr.constr) (y : EConstr.constr) + : (EConstr.constr -> EConstr.constr) option = + let open EConstr in + let open Vars in let open Context.Rel.Declaration in let rec coerce_unify env x y = - let x = hnf env !evdref (EConstr.of_constr x) and y = hnf env !evdref (EConstr.of_constr y) in + let x = hnf env !evdref x and y = hnf env !evdref y in + let x = EConstr.of_constr x in + let y = EConstr.of_constr y in try - evdref := the_conv_x_leq env (EConstr.of_constr x) (EConstr.of_constr y) !evdref; + evdref := the_conv_x_leq env x y !evdref; None with UnableToUnify _ -> coerce' env x y - and coerce' env x y : (Term.constr -> Term.constr) option = + and coerce' env x y : (EConstr.constr -> EConstr.constr) option = let subco () = subset_coerce env evdref x y in let dest_prod c = - match Reductionops.splay_prod_n env (!evdref) 1 (EConstr.of_constr c) with - | [LocalAssum (na,t) | LocalDef (na,_,t)], c -> (na,t), c + match Reductionops.splay_prod_n env (!evdref) 1 c with + | [LocalAssum (na,t) | LocalDef (na,_,t)], c -> (na, EConstr.of_constr t), EConstr.of_constr c | _ -> raise NoSubtacCoercion in let coerce_application typ typ' c c' l l' = @@ -162,7 +173,7 @@ and coerce loc env evdref (x : Term.constr) (y : Term.constr) let rec aux tele typ typ' i co = if i < len then let hdx = l.(i) and hdy = l'.(i) in - try evdref := the_conv_x_leq env (EConstr.of_constr hdx) (EConstr.of_constr hdy) !evdref; + try evdref := the_conv_x_leq env hdx hdy !evdref; let (n, eqT), restT = dest_prod typ in let (n', eqT'), restT' = dest_prod typ' in aux (hdx :: tele) (subst1 hdx restT) (subst1 hdy restT') (succ i) co @@ -170,16 +181,16 @@ and coerce loc env evdref (x : Term.constr) (y : Term.constr) let (n, eqT), restT = dest_prod typ in let (n', eqT'), restT' = dest_prod typ' in let _ = - try evdref := the_conv_x_leq env (EConstr.of_constr eqT) (EConstr.of_constr eqT') !evdref + try evdref := the_conv_x_leq env eqT eqT' !evdref with UnableToUnify _ -> raise NoSubtacCoercion in (* Disallow equalities on arities *) - if Reduction.is_arity env eqT then raise NoSubtacCoercion; + if Reductionops.is_arity env !evdref eqT then raise NoSubtacCoercion; let restargs = lift_args 1 (List.rev (Array.to_list (Array.sub l (succ i) (len - (succ i))))) in let args = List.rev (restargs @ mkRel 1 :: List.map (lift 1) tele) in - let pred = mkLambda (n, eqT, applistc (lift 1 c) args) in + let pred = mkLambda (n, eqT, applist (lift 1 c, args)) in let eq = papp evdref coq_eq_ind [| eqT; hdx; hdy |] in let evar = make_existential loc env evdref eq in let eq_app x = papp evdref coq_eq_rect @@ -188,15 +199,15 @@ and coerce loc env evdref (x : Term.constr) (y : Term.constr) aux (hdy :: tele) (subst1 hdx restT) (subst1 hdy restT') (succ i) (fun x -> eq_app (co x)) else Some (fun x -> - let term = EConstr.of_constr (co x) in - Typing.e_solve_evars env evdref term) + let term = co x in + EConstr.of_constr (Typing.e_solve_evars env evdref term)) in - if isEvar c || isEvar c' || not (Program.is_program_generalized_coercion ()) then + if isEvar !evdref c || isEvar !evdref c' || not (Program.is_program_generalized_coercion ()) then (* Second-order unification needed. *) raise NoSubtacCoercion; aux [] typ typ' 0 (fun x -> x) in - match (kind_of_term x, kind_of_term y) with + match (EConstr.kind !evdref x, EConstr.kind !evdref y) with | Sort s, Sort s' -> (match s, s' with | Prop x, Prop y when x == y -> None @@ -207,7 +218,7 @@ and coerce loc env evdref (x : Term.constr) (y : Term.constr) let name' = Name (Namegen.next_ident_away Namegen.default_dependent_ident (Termops.ids_of_context env)) in - let env' = push_rel (LocalAssum (name', a')) env in + let env' = push_rel (local_assum (name', a')) env in let c1 = coerce_unify env' (lift 1 a') (lift 1 a) in (* env, x : a' |- c1 : lift 1 a' > lift 1 a *) let coec1 = app_opt env' evdref c1 (mkRel 1) in @@ -224,7 +235,7 @@ and coerce loc env evdref (x : Term.constr) (y : Term.constr) (mkApp (lift 1 f, [| coec1 |]))))) | App (c, l), App (c', l') -> - (match kind_of_term c, kind_of_term c' with + (match EConstr.kind !evdref c, EConstr.kind !evdref c' with Ind (i, u), Ind (i', u') -> (* Inductive types *) let len = Array.length l in let sigT = delayed_force sigT_typ in @@ -241,23 +252,21 @@ and coerce loc env evdref (x : Term.constr) (y : Term.constr) in let c1 = coerce_unify env a a' in let remove_head a c = - match kind_of_term c with + match EConstr.kind !evdref c with | Lambda (n, t, t') -> c, t' | Evar (k, args) -> - let (evs, t) = Evardefine.define_evar_as_lambda env !evdref (k, Array.map EConstr.of_constr args) in + let (evs, t) = Evardefine.define_evar_as_lambda env !evdref (k,args) in evdref := evs; - let t = EConstr.Unsafe.to_constr t in - let (n, dom, rng) = destLambda t in - let dom = whd_evar !evdref dom in - if isEvar dom then - let (domk, args) = destEvar dom in - evdref := define domk a !evdref; + let (n, dom, rng) = destLambda !evdref t in + if isEvar !evdref dom then + let (domk, args) = destEvar !evdref dom in + evdref := define domk (EConstr.Unsafe.to_constr a) !evdref; else (); t, rng | _ -> raise NoSubtacCoercion in let (pb, b), (pb', b') = remove_head a pb, remove_head a' pb' in - let env' = push_rel (LocalAssum (Name Namegen.default_dependent_ident, a)) env in + let env' = push_rel (local_assum (Name Namegen.default_dependent_ident, a)) env in let c2 = coerce_unify env' b b' in match c1, c2 with | None, None -> None @@ -297,30 +306,30 @@ and coerce loc env evdref (x : Term.constr) (y : Term.constr) let evm = !evdref in (try subco () with NoSubtacCoercion -> - let typ = Typing.unsafe_type_of env evm (EConstr.of_constr c) in - let typ' = Typing.unsafe_type_of env evm (EConstr.of_constr c') in - coerce_application typ typ' c c' l l') + let typ = Typing.unsafe_type_of env evm c in + let typ' = Typing.unsafe_type_of env evm c' in + coerce_application (EConstr.of_constr typ) (EConstr.of_constr typ') c c' l l') else subco () - | x, y when Constr.equal c c' -> + | x, y when EConstr.eq_constr !evdref c c' -> if Int.equal (Array.length l) (Array.length l') then let evm = !evdref in - let lam_type = Typing.unsafe_type_of env evm (EConstr.of_constr c) in - let lam_type' = Typing.unsafe_type_of env evm (EConstr.of_constr c') in - coerce_application lam_type lam_type' c c' l l' + let lam_type = Typing.unsafe_type_of env evm c in + let lam_type' = Typing.unsafe_type_of env evm c' in + coerce_application (EConstr.of_constr lam_type) (EConstr.of_constr lam_type') c c' l l' else subco () | _ -> subco ()) | _, _ -> subco () and subset_coerce env evdref x y = - match disc_subset x with + match disc_subset !evdref x with Some (u, p) -> let c = coerce_unify env u y in let f x = app_opt env evdref c (papp evdref sig_proj1 [| u; p; x |]) in Some f | None -> - match disc_subset y with + match disc_subset !evdref y with Some (u, p) -> let c = coerce_unify env x u in Some @@ -337,8 +346,8 @@ let app_coercion env evdref coercion v = match coercion with | None -> v | Some f -> - let v' = Typing.e_solve_evars env evdref (EConstr.of_constr (f v)) in - whd_betaiota !evdref (EConstr.of_constr v') + let v' = Typing.e_solve_evars env evdref (f v) in + EConstr.of_constr (whd_betaiota !evdref (EConstr.of_constr v')) let coerce_itf loc env evd v t c1 = let evdref = ref evd in @@ -358,7 +367,7 @@ let apply_coercion env sigma p hj typ_cl = (fun (ja,typ_cl,sigma) i -> let ((fv,isid,isproj),ctx) = coercion_value i in let sigma = Evd.merge_context_set Evd.univ_flexible sigma ctx in - let argl = (class_args_of env sigma (EConstr.of_constr typ_cl))@[ja.uj_val] in + let argl = (class_args_of env sigma typ_cl)@[EConstr.of_constr ja.uj_val] in let sigma, jres = apply_coercion_args env sigma true isproj argl fv in @@ -366,7 +375,7 @@ let apply_coercion env sigma p hj typ_cl = { uj_val = ja.uj_val; uj_type = jres.uj_type } else jres), - jres.uj_type,sigma) + EConstr.of_constr jres.uj_type,sigma) (hj,typ_cl,sigma) p in evd, j with NoCoercion as e -> raise e @@ -375,7 +384,8 @@ let apply_coercion env sigma p hj typ_cl = (* Try to coerce to a funclass; raise NoCoercion if not possible *) let inh_app_fun_core env evd j = let t = whd_all env evd (EConstr.of_constr j.uj_type) in - match EConstr.kind evd (EConstr.of_constr t) with + let t = EConstr.of_constr t in + match EConstr.kind evd t with | Prod (_,_,_) -> (evd,j) | Evar ev -> let (evd',t) = Evardefine.define_evar_as_product evd ev in @@ -389,7 +399,7 @@ let inh_app_fun_core env evd j = try let evdref = ref evd in let coercef, t = mu env evdref t in - let res = { uj_val = app_opt env evdref coercef j.uj_val; uj_type = t } in + let res = { uj_val = EConstr.Unsafe.to_constr (app_opt env evdref coercef (EConstr.of_constr j.uj_val)); uj_type = EConstr.Unsafe.to_constr t } in (!evdref, res) with NoSubtacCoercion | NoCoercion -> (evd,j) @@ -427,10 +437,10 @@ let inh_coerce_to_sort loc env evd j = let inh_coerce_to_base loc env evd j = if Flags.is_program_mode () then let evdref = ref evd in - let ct, typ' = mu env evdref j.uj_type in + let ct, typ' = mu env evdref (EConstr.of_constr j.uj_type) in let res = - { uj_val = app_coercion env evdref ct j.uj_val; - uj_type = typ' } + { uj_val = EConstr.Unsafe.to_constr (app_coercion env evdref ct (EConstr.of_constr j.uj_val)); + uj_type = EConstr.Unsafe.to_constr typ' } in !evdref, res else (evd, j) @@ -442,33 +452,35 @@ let inh_coerce_to_prod loc env evd t = else (evd, t) let inh_coerce_to_fail env evd rigidonly v t c1 = - if rigidonly && not (Heads.is_rigid env c1 && Heads.is_rigid env t) + if rigidonly && not (Heads.is_rigid env (EConstr.Unsafe.to_constr c1) && Heads.is_rigid env (EConstr.Unsafe.to_constr t)) then raise NoCoercion else let evd, v', t' = try - let t2,t1,p = lookup_path_between env evd (EConstr.of_constr t,EConstr.of_constr c1) in + let t2,t1,p = lookup_path_between env evd (t,c1) in match v with | Some v -> let evd,j = apply_coercion env evd p - {uj_val = v; uj_type = t} t2 in - evd, Some j.uj_val, j.uj_type + {uj_val = EConstr.Unsafe.to_constr v; uj_type = EConstr.Unsafe.to_constr t} t2 in + evd, Some (EConstr.of_constr j.uj_val), (EConstr.of_constr j.uj_type) | None -> evd, None, t with Not_found -> raise NoCoercion in - try (the_conv_x_leq env (EConstr.of_constr t') (EConstr.of_constr c1) evd, v') + try (the_conv_x_leq env t' c1 evd, v') with UnableToUnify _ -> raise NoCoercion let rec inh_conv_coerce_to_fail loc env evd rigidonly v t c1 = - try (the_conv_x_leq env (EConstr.of_constr t) (EConstr.of_constr c1) evd, v) + let open EConstr in + let open Vars in + try (the_conv_x_leq env t c1 evd, v) with UnableToUnify (best_failed_evd,e) -> try inh_coerce_to_fail env evd rigidonly v t c1 with NoCoercion -> match - kind_of_term (whd_all env evd (EConstr.of_constr t)), - kind_of_term (whd_all env evd (EConstr.of_constr c1)) + EConstr.kind evd (EConstr.of_constr (whd_all env evd t)), + EConstr.kind evd (EConstr.of_constr (whd_all env evd c1)) with | Prod (name,t1,t2), Prod (_,u1,u2) -> (* Conversion did not work, we may succeed with a coercion. *) @@ -481,16 +493,16 @@ let rec inh_conv_coerce_to_fail loc env evd rigidonly v t c1 = | Anonymous -> Name Namegen.default_dependent_ident | _ -> name in let open Context.Rel.Declaration in - let env1 = push_rel (LocalAssum (name,u1)) env in + let env1 = push_rel (local_assum (name,u1)) env in let (evd', v1) = inh_conv_coerce_to_fail loc env1 evd rigidonly (Some (mkRel 1)) (lift 1 u1) (lift 1 t1) in let v1 = Option.get v1 in - let v2 = Option.map (fun v -> beta_applist evd' (EConstr.of_constr (lift 1 v),[EConstr.of_constr v1])) v in + let v2 = Option.map (fun v -> beta_applist evd' (lift 1 v,[v1])) v in let t2 = match v2 with - | None -> subst_term evd' (EConstr.of_constr v1) (EConstr.of_constr t2) + | None -> subst_term evd' v1 t2 | Some v2 -> Retyping.get_type_of env1 evd' (EConstr.of_constr v2) in - let (evd'',v2') = inh_conv_coerce_to_fail loc env1 evd' rigidonly v2 t2 u2 in + let (evd'',v2') = inh_conv_coerce_to_fail loc env1 evd' rigidonly (Option.map EConstr.of_constr v2) (EConstr.of_constr t2) u2 in (evd'', Option.map (fun v2' -> mkLambda (name, u1, v2')) v2') | _ -> raise (NoCoercionNoUnifier (best_failed_evd,e)) @@ -498,27 +510,27 @@ let rec inh_conv_coerce_to_fail loc env evd rigidonly v t c1 = let inh_conv_coerce_to_gen resolve_tc rigidonly loc env evd cj t = let (evd', val') = try - inh_conv_coerce_to_fail loc env evd rigidonly (Some cj.uj_val) cj.uj_type t + inh_conv_coerce_to_fail loc env evd rigidonly (Some (EConstr.of_constr cj.uj_val)) (EConstr.of_constr cj.uj_type) t with NoCoercionNoUnifier (best_failed_evd,e) -> try if Flags.is_program_mode () then - coerce_itf loc env evd (Some cj.uj_val) cj.uj_type t + coerce_itf loc env evd (Some (EConstr.of_constr cj.uj_val)) (EConstr.of_constr cj.uj_type) t else raise NoSubtacCoercion with | NoSubtacCoercion when not resolve_tc || not !use_typeclasses_for_conversion -> - error_actual_type ~loc env best_failed_evd cj t e + error_actual_type ~loc env best_failed_evd cj (EConstr.Unsafe.to_constr t) e | NoSubtacCoercion -> let evd' = saturate_evd env evd in try if evd' == evd then - error_actual_type ~loc env best_failed_evd cj t e + error_actual_type ~loc env best_failed_evd cj (EConstr.Unsafe.to_constr t) e else - inh_conv_coerce_to_fail loc env evd' rigidonly (Some cj.uj_val) cj.uj_type t + inh_conv_coerce_to_fail loc env evd' rigidonly (Some (EConstr.of_constr cj.uj_val)) (EConstr.of_constr cj.uj_type) t with NoCoercionNoUnifier (_evd,_error) -> - error_actual_type ~loc env best_failed_evd cj t e + error_actual_type ~loc env best_failed_evd cj (EConstr.Unsafe.to_constr t) e in let val' = match val' with Some v -> v | None -> assert(false) in - (evd',{ uj_val = val'; uj_type = t }) + (evd',{ uj_val = EConstr.Unsafe.to_constr val'; uj_type = EConstr.Unsafe.to_constr t }) let inh_conv_coerce_to resolve_tc = inh_conv_coerce_to_gen resolve_tc false let inh_conv_coerce_rigid_to resolve_tc = inh_conv_coerce_to_gen resolve_tc true diff --git a/pretyping/coercion.mli b/pretyping/coercion.mli index 68f9a2e68..62d4fb004 100644 --- a/pretyping/coercion.mli +++ b/pretyping/coercion.mli @@ -36,7 +36,7 @@ val inh_coerce_to_base : Loc.t -> (** [inh_coerce_to_prod env isevars t] coerces [t] to a product type *) val inh_coerce_to_prod : Loc.t -> - env -> evar_map -> types -> evar_map * types + env -> evar_map -> EConstr.types -> evar_map * EConstr.types (** [inh_conv_coerce_to resolve_tc Loc.t env isevars j t] coerces [j] to an object of type [t]; i.e. it inserts a coercion into [j], if needed, in such @@ -44,16 +44,16 @@ val inh_coerce_to_prod : Loc.t -> applicable. resolve_tc=false disables resolving type classes (as the last resort before failing) *) val inh_conv_coerce_to : bool -> Loc.t -> - env -> evar_map -> unsafe_judgment -> types -> evar_map * unsafe_judgment + env -> evar_map -> unsafe_judgment -> EConstr.types -> evar_map * unsafe_judgment val inh_conv_coerce_rigid_to : bool -> Loc.t -> - env -> evar_map -> unsafe_judgment -> types -> evar_map * unsafe_judgment + env -> evar_map -> unsafe_judgment -> EConstr.types -> evar_map * unsafe_judgment (** [inh_conv_coerces_to loc env isevars t t'] checks if an object of type [t] is coercible to an object of type [t'] adding evar constraints if needed; it fails if no coercion exists *) val inh_conv_coerces_to : Loc.t -> - env -> evar_map -> types -> types -> evar_map + env -> evar_map -> EConstr.types -> EConstr.types -> evar_map (** [inh_pattern_coerce_to loc env isevars pat ind1 ind2] coerces the Cases pattern [pat] typed in [ind1] into a pattern typed in [ind2]; diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml index 28ba60812..18731f1e9 100644 --- a/pretyping/pretyping.ml +++ b/pretyping/pretyping.ml @@ -359,7 +359,7 @@ let allow_anonymous_refs = ref false let inh_conv_coerce_to_tycon resolve_tc loc env evdref j = function | None -> j | Some t -> - evd_comb2 (Coercion.inh_conv_coerce_to resolve_tc loc env.ExtraEnv.env) evdref j t + evd_comb2 (Coercion.inh_conv_coerce_to resolve_tc loc env.ExtraEnv.env) evdref j (EConstr.of_constr t) let check_instance loc subst = function | [] -> () @@ -770,8 +770,8 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre match tycon with | None -> evd, tycon | Some ty -> - let evd, ty' = Coercion.inh_coerce_to_prod loc env.ExtraEnv.env evd ty in - evd, Some ty') + let evd, ty' = Coercion.inh_coerce_to_prod loc env.ExtraEnv.env evd (EConstr.of_constr ty) in + evd, Some (EConstr.Unsafe.to_constr ty')) evdref tycon in let (name',dom,rng) = evd_comb1 (split_tycon loc env.ExtraEnv.env) evdref tycon' in diff --git a/pretyping/program.ml b/pretyping/program.ml index 4b6137b53..2606d91f3 100644 --- a/pretyping/program.ml +++ b/pretyping/program.ml @@ -29,8 +29,9 @@ let init_constant dir s () = coq_constant "Program" dir s let init_reference dir s () = coq_reference "Program" dir s let papp evdref r args = + let open EConstr in let gr = delayed_force r in - mkApp (Evarutil.e_new_global evdref gr, args) + mkApp (EConstr.of_constr (Evarutil.e_new_global evdref gr), args) let sig_typ = init_reference ["Init"; "Specif"] "sig" let sig_intro = init_reference ["Init"; "Specif"] "exist" diff --git a/pretyping/program.mli b/pretyping/program.mli index 023ff8ca5..64c4ca2c2 100644 --- a/pretyping/program.mli +++ b/pretyping/program.mli @@ -36,7 +36,7 @@ val mk_coq_and : constr list -> constr val mk_coq_not : constr -> constr (** Polymorphic application of delayed references *) -val papp : Evd.evar_map ref -> (unit -> global_reference) -> constr array -> constr +val papp : Evd.evar_map ref -> (unit -> global_reference) -> EConstr.constr array -> EConstr.constr val get_proofs_transparency : unit -> bool val is_program_cases : unit -> bool diff --git a/pretyping/unification.ml b/pretyping/unification.ml index 786cfd31f..b568dd044 100644 --- a/pretyping/unification.ml +++ b/pretyping/unification.ml @@ -1225,7 +1225,7 @@ let is_mimick_head ts f = let try_to_coerce env evd c cty tycon = let j = make_judge c cty in - let (evd',j') = inh_conv_coerce_rigid_to true Loc.ghost env evd j tycon in + let (evd',j') = inh_conv_coerce_rigid_to true Loc.ghost env evd j (EConstr.of_constr tycon) in let evd' = Evarconv.consider_remaining_unif_problems env evd' in let evd' = Evd.map_metas_fvalue (nf_evar evd') evd' in (evd',j'.uj_val) diff --git a/proofs/clenv.ml b/proofs/clenv.ml index c2130a64a..0515e4198 100644 --- a/proofs/clenv.ml +++ b/proofs/clenv.ml @@ -670,7 +670,7 @@ let define_with_type sigma env ev c = let t = Retyping.get_type_of env sigma (EConstr.of_constr ev) in let ty = Retyping.get_type_of env sigma (EConstr.of_constr c) in let j = Environ.make_judge c ty in - let (sigma, j) = Coercion.inh_conv_coerce_to true (Loc.ghost) env sigma j t in + let (sigma, j) = Coercion.inh_conv_coerce_to true (Loc.ghost) env sigma j (EConstr.of_constr t) in let (ev, _) = destEvar ev in let sigma = Evd.define ev j.Environ.uj_val sigma in sigma diff --git a/proofs/refine.ml b/proofs/refine.ml index b62f0bea4..19134bfa3 100644 --- a/proofs/refine.ml +++ b/proofs/refine.ml @@ -136,7 +136,7 @@ let with_type env evd c t = let my_type = Retyping.get_type_of env evd (EConstr.of_constr c) in let j = Environ.make_judge c my_type in let (evd,j') = - Coercion.inh_conv_coerce_to true (Loc.ghost) env evd j t + Coercion.inh_conv_coerce_to true (Loc.ghost) env evd j (EConstr.of_constr t) in evd , j'.Environ.uj_val diff --git a/toplevel/class.ml b/toplevel/class.ml index 6295eb336..46b212dee 100644 --- a/toplevel/class.ml +++ b/toplevel/class.ml @@ -85,7 +85,7 @@ let uniform_cond nargs lt = let rec aux = function | (0,[]) -> true | (n,t::l) -> - let t = strip_outer_cast Evd.empty (EConstr.of_constr t) (** FIXME *) in + let t = strip_outer_cast Evd.empty t (** FIXME *) in isRel t && Int.equal (destRel t) n && aux ((n-1),l) | _ -> false in -- cgit v1.2.3 From 67dc22d8389234d0c9b329944ff579e7056b7250 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 8 Nov 2016 10:57:05 +0100 Subject: Cases API using EConstr. --- engine/termops.ml | 14 +- engine/termops.mli | 6 +- ltac/rewrite.ml | 2 +- plugins/decl_mode/decl_proof_instr.ml | 8 +- plugins/firstorder/formula.ml | 2 +- plugins/funind/functional_principles_proofs.ml | 2 +- pretyping/cases.ml | 454 ++++++++++++++----------- pretyping/cases.mli | 11 +- pretyping/indrec.ml | 5 +- pretyping/inductiveops.ml | 13 +- pretyping/inductiveops.mli | 4 +- pretyping/program.ml | 6 +- pretyping/program.mli | 4 +- tactics/equality.ml | 2 +- tactics/hipattern.ml | 4 +- toplevel/class.ml | 2 +- toplevel/himsg.ml | 5 +- toplevel/indschemes.ml | 2 +- 18 files changed, 314 insertions(+), 232 deletions(-) diff --git a/engine/termops.ml b/engine/termops.ml index e5db3c085..16e2c04c8 100644 --- a/engine/termops.ml +++ b/engine/termops.ml @@ -207,10 +207,11 @@ let mkProd_or_LetIn decl c = (* 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 + | LocalAssum (na,t) -> mkProd (na, EConstr.of_constr t, c) + | LocalDef (_,b,_) -> Vars.subst1 (EConstr.of_constr 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 @@ -1099,6 +1100,15 @@ let global_app_of_constr sigma c = | Proj (p, c) -> (ConstRef (Projection.constant p), Instance.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 } diff --git a/engine/termops.mli b/engine/termops.mli index c90fdf9c2..2b66c88a7 100644 --- a/engine/termops.mli +++ b/engine/termops.mli @@ -45,11 +45,11 @@ 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_wo_LetIn : Context.Rel.Declaration.t -> EConstr.types -> EConstr.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_mkProd_wo_LetIn : EConstr.types -> Context.Rel.t -> EConstr.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 @@ -170,6 +170,8 @@ val eta_reduce_head : constr -> constr (** Flattens application lists *) val collapse_appl : Evd.evar_map -> EConstr.t -> constr +val prod_applist : Evd.evar_map -> EConstr.t -> EConstr.t list -> EConstr.t + (** 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 -> EConstr.t -> constr diff --git a/ltac/rewrite.ml b/ltac/rewrite.ml index 0091d0d0a..58153c453 100644 --- a/ltac/rewrite.ml +++ b/ltac/rewrite.ml @@ -344,7 +344,7 @@ end) = struct Some (evars, found, c, ty, arg :: args) with Not_found -> let ty = whd_all env ty in - find env (mkApp (c, [| arg |])) (prod_applist ty [arg]) args + find env (mkApp (c, [| arg |])) (Term.prod_applist ty [arg]) args in find env c ty args let unlift_cstr env sigma = function diff --git a/plugins/decl_mode/decl_proof_instr.ml b/plugins/decl_mode/decl_proof_instr.ml index dcebf7806..e587fd52c 100644 --- a/plugins/decl_mode/decl_proof_instr.ml +++ b/plugins/decl_mode/decl_proof_instr.ml @@ -352,7 +352,7 @@ let enstack_subsubgoals env se stack gls= let constructor = mkConstructU ((ind,succ i),u) (* constructors numbering*) in let appterm = applist (constructor,params) in - let apptype = prod_applist gentyp params in + let apptype = Term.prod_applist gentyp params in let rc,_ = Reduction.dest_prod env apptype in let rec meta_aux last lenv = function [] -> (last,lenv,[]) @@ -687,7 +687,7 @@ let rec build_applist prod = function [] -> [],prod | n::q -> let (_,typ,_) = destProd prod in - let ctx,head = build_applist (prod_applist prod [mkMeta n]) q in + let ctx,head = build_applist (Term.prod_applist prod [mkMeta n]) q in (n,typ)::ctx,head let instr_suffices _then cut gls0 = @@ -720,7 +720,7 @@ let conjunction_arity id gls = let gentypes= Inductive.arities_of_constructors indu (mib,oib) in let _ = if not (Int.equal (Array.length gentypes) 1) then raise Not_found in - let apptype = prod_applist gentypes.(0) params in + let apptype = Term.prod_applist gentypes.(0) params in let rc,_ = Reduction.dest_prod env apptype in List.length rc | _ -> raise Not_found @@ -1270,7 +1270,7 @@ let rec execute_cases fix_name per_info tacnext args objs nhrec tree gls = let gen_arities = Inductive.arities_of_constructors (ind,u) spec in let f_ids typ = let sign = - (prod_assum (prod_applist typ params)) in + (prod_assum (Term.prod_applist typ params)) in find_intro_names sign gls in let constr_args_ids = Array.map f_ids gen_arities in let case_term = diff --git a/plugins/firstorder/formula.ml b/plugins/firstorder/formula.ml index 79f185d18..60e9196af 100644 --- a/plugins/firstorder/formula.ml +++ b/plugins/firstorder/formula.ml @@ -54,7 +54,7 @@ let construct_nhyps ind gls = let ind_hyps nevar ind largs gls= let types= Inductiveops.arities_of_constructors (pf_env gls) ind in let myhyps t = - let t1=prod_applist t largs in + let t1=Term.prod_applist t largs in let t2=snd (decompose_prod_n_assum nevar t1) in fst (decompose_prod_assum t2) in Array.map myhyps types diff --git a/plugins/funind/functional_principles_proofs.ml b/plugins/funind/functional_principles_proofs.ml index 0a7938069..83fc48623 100644 --- a/plugins/funind/functional_principles_proofs.ml +++ b/plugins/funind/functional_principles_proofs.ml @@ -1150,7 +1150,7 @@ let prove_princ_for_struct (evd:Evd.evar_map ref) interactive_proof fun_num fnam let info_array = Array.mapi (fun i types -> - let types = prod_applist types (List.rev_map var_of_decl princ_params) in + let types = Term.prod_applist types (List.rev_map var_of_decl princ_params) in { idx = idxs.(i) - fix_offset; name = Nameops.out_name (fresh_id names.(i)); types = types; diff --git a/pretyping/cases.ml b/pretyping/cases.ml index 96c61647c..1a181202c 100644 --- a/pretyping/cases.ml +++ b/pretyping/cases.ml @@ -6,14 +6,17 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +module CVars = Vars + open Pp open CErrors open Util open Names open Nameops open Term -open Vars open Termops +open EConstr +open Vars open Namegen open Declarations open Inductiveops @@ -35,6 +38,14 @@ open Context.Rel.Declaration module RelDecl = Context.Rel.Declaration module NamedDecl = Context.Named.Declaration +let local_assum (na, t) = + let inj = EConstr.Unsafe.to_constr in + RelDecl.LocalAssum (na, inj t) + +let local_def (na, b, t) = + let inj = EConstr.Unsafe.to_constr in + RelDecl.LocalDef (na, inj b, inj t) + (* Pattern-matching errors *) type pattern_matching_error = @@ -78,6 +89,9 @@ let list_try_compile f l = let force_name = let nx = Name default_dependent_ident in function Anonymous -> nx | na -> na +let make_judge c ty = + make_judge (EConstr.Unsafe.to_constr c) (EConstr.Unsafe.to_constr ty) + (************************************************************************) (* Pattern-matching compilation (Cases) *) (************************************************************************) @@ -99,11 +113,13 @@ let make_anonymous_patvars n = let relocate_rel n1 n2 k j = if Int.equal j (n1 + k) then n2+k else j -let rec relocate_index n1 n2 k t = match kind_of_term t with +let rec relocate_index sigma n1 n2 k t = + let open EConstr in + match EConstr.kind sigma t with | Rel j when Int.equal j (n1 + k) -> mkRel (n2+k) | Rel j when j < n1+k -> t | Rel j when j > n1+k -> t - | _ -> map_constr_with_binders succ (relocate_index n1 n2) k t + | _ -> EConstr.map_with_binders sigma succ (relocate_index sigma n1 n2) k t (**********************************************************************) (* Structures used in compiling pattern-matching *) @@ -283,16 +299,18 @@ let inductive_template evdref env tmloc ind = (fun decl (subst,evarl,n) -> match decl with | LocalAssum (na,ty) -> + let ty = EConstr.of_constr ty in let ty' = substl subst ty in - let e = e_new_evar env evdref ~src:(hole_source n) ty' in + let e = EConstr.of_constr (e_new_evar env evdref ~src:(hole_source n) (EConstr.Unsafe.to_constr ty')) in (e::subst,e::evarl,n+1) | LocalDef (na,b,ty) -> + let b = EConstr.of_constr b in (substl subst b::subst,evarl,n+1)) arsign ([],[],1) in applist (mkIndU indu,List.rev evarl) let try_find_ind env sigma typ realnames = - let (IndType(indf,realargs) as ind) = find_rectype env sigma (EConstr.of_constr typ) in + let (IndType(indf,realargs) as ind) = find_rectype env sigma typ in let names = match realnames with | Some names -> names @@ -308,21 +326,23 @@ let inh_coerce_to_ind evdref env loc ty tyi = constructor and renounce if not able to give more information *) (* devrait être indifférent d'exiger leq ou pas puisque pour un inductif cela doit être égal *) - if not (e_cumul env evdref (EConstr.of_constr expected_typ) (EConstr.of_constr ty)) then evdref := sigma + if not (e_cumul env evdref expected_typ ty) then evdref := sigma -let binding_vars_of_inductive = function +let binding_vars_of_inductive sigma = function | NotInd _ -> [] - | IsInd (_,IndType(_,realargs),_) -> List.filter isRel realargs + | IsInd (_,IndType(_,realargs),_) -> List.filter (isRel sigma) (List.map EConstr.of_constr realargs) let extract_inductive_data env sigma decl = match decl with | LocalAssum (_,t) -> + let t = EConstr.of_constr t in let tmtyp = try try_find_ind env sigma t None with Not_found -> NotInd (None,t) in - let tmtypvars = binding_vars_of_inductive tmtyp in + let tmtypvars = binding_vars_of_inductive sigma tmtyp in (tmtyp,tmtypvars) | LocalDef (_,_,t) -> + let t = EConstr.of_constr t in (NotInd (None, t), []) let unify_tomatch_with_patterns evdref env loc typ pats realnames = @@ -336,7 +356,7 @@ let unify_tomatch_with_patterns evdref env loc typ pats realnames = let find_tomatch_tycon evdref env loc = function (* Try if some 'in I ...' is present and can be used as a constraint *) | Some (_,ind,realnal) -> - mk_tycon (EConstr.of_constr (inductive_template evdref env loc ind)),Some (List.rev realnal) + mk_tycon (inductive_template evdref env loc ind),Some (List.rev realnal) | None -> empty_tycon,None @@ -346,12 +366,12 @@ let coerce_row typing_fun evdref env pats (tomatch,(_,indopt)) = let j = typing_fun tycon env evdref tomatch in let evd, j = Coercion.inh_coerce_to_base (loc_of_glob_constr tomatch) env !evdref j in evdref := evd; - let typ = nf_evar !evdref j.uj_type in + let typ = EConstr.of_constr (nf_evar !evdref j.uj_type) in let t = try try_find_ind env !evdref typ realnames with Not_found -> unify_tomatch_with_patterns evdref env loc typ pats realnames in - (j.uj_val,t) + (EConstr.of_constr j.uj_val,t) let coerce_to_indtype typing_fun evdref env matx tomatchl = let pats = List.map (fun r -> r.patterns) matx in @@ -364,7 +384,7 @@ let coerce_to_indtype typing_fun evdref env matx tomatchl = (* Utils *) let mkExistential env ?(src=(Loc.ghost,Evar_kinds.InternalHole)) evdref = - let e, u = e_new_type_evar env evdref univ_flexible_alg ~src:src in e + let e, u = e_new_type_evar env evdref univ_flexible_alg ~src:src in EConstr.of_constr e let evd_comb2 f evdref x y = let (evd',y) = f !evdref x y in @@ -390,13 +410,13 @@ let adjust_tomatch_to_pattern pb ((current,typ),deps,dep) = | Some (_,(ind,_)) -> let indt = inductive_template pb.evdref pb.env None ind in let current = - if List.is_empty deps && isEvar typ then + if List.is_empty deps && isEvar !(pb.evdref) typ then (* Don't insert coercions if dependent; only solve evars *) - let _ = e_cumul pb.env pb.evdref (EConstr.of_constr indt) (EConstr.of_constr typ) in + let _ = e_cumul pb.env pb.evdref indt typ in current else - (evd_comb2 (Coercion.inh_conv_coerce_to true Loc.ghost pb.env) - pb.evdref (make_judge current typ) (EConstr.of_constr indt)).uj_val in + EConstr.of_constr (evd_comb2 (Coercion.inh_conv_coerce_to true Loc.ghost pb.env) + pb.evdref (make_judge current typ) indt).uj_val in let sigma = !(pb.evdref) in (current,try_find_ind pb.env sigma indt names)) | _ -> (current,tmtyp) @@ -406,10 +426,10 @@ let type_of_tomatch = function | NotInd (_,t) -> t let map_tomatch_type f = function - | IsInd (t,ind,names) -> IsInd (f t,map_inductive_type f ind,names) + | IsInd (t,ind,names) -> IsInd (f t,map_inductive_type (fun c -> EConstr.Unsafe.to_constr (f (EConstr.of_constr c))) ind,names) | NotInd (c,t) -> NotInd (Option.map f c, f t) -let liftn_tomatch_type n depth = map_tomatch_type (liftn n depth) +let liftn_tomatch_type n depth = map_tomatch_type (Vars.liftn n depth) let lift_tomatch_type n = liftn_tomatch_type n 1 (**********************************************************************) @@ -435,7 +455,7 @@ let remove_current_pattern eqn = let push_current_pattern (cur,ty) eqn = match eqn.patterns with | pat::pats -> - let rhs_env = push_rel (LocalDef (alias_of_pat pat,cur,ty)) eqn.rhs.rhs_env in + let rhs_env = push_rel (local_def (alias_of_pat pat,cur,ty)) eqn.rhs.rhs_env in { eqn with rhs = { eqn.rhs with rhs_env = rhs_env }; patterns = pats } @@ -537,8 +557,8 @@ let dependencies_in_pure_rhs nargs eqns = let dependent_decl sigma a = function - | LocalAssum (na,t) -> dependent sigma (EConstr.of_constr a) (EConstr.of_constr t) - | LocalDef (na,c,t) -> dependent sigma (EConstr.of_constr a) (EConstr.of_constr t) || dependent sigma (EConstr.of_constr a) (EConstr.of_constr c) + | LocalAssum (na,t) -> dependent sigma a (EConstr.of_constr t) + | LocalDef (na,c,t) -> dependent sigma a (EConstr.of_constr t) || dependent sigma a (EConstr.of_constr c) let rec dep_in_tomatch sigma n = function | (Pushed _ | Alias _ | NonDepAlias) :: l -> dep_in_tomatch sigma n l @@ -546,7 +566,7 @@ let rec dep_in_tomatch sigma n = function | [] -> false let dependencies_in_rhs sigma nargs current tms eqns = - match kind_of_term current with + match EConstr.kind sigma current with | Rel n when dep_in_tomatch sigma n tms -> List.make nargs true | _ -> dependencies_in_pure_rhs nargs eqns @@ -593,31 +613,31 @@ let find_dependencies_signature sigma deps_in_rhs typs = [relocate_index_tomatch 1 n tomatch] will go the way back. *) -let relocate_index_tomatch n1 n2 = +let relocate_index_tomatch sigma n1 n2 = let rec genrec depth = function | [] -> [] | Pushed (b,((c,tm),l,na)) :: rest -> - let c = relocate_index n1 n2 depth c in - let tm = map_tomatch_type (relocate_index n1 n2 depth) tm in + let c = relocate_index sigma n1 n2 depth c in + let tm = map_tomatch_type (relocate_index sigma n1 n2 depth) tm in let l = List.map (relocate_rel n1 n2 depth) l in Pushed (b,((c,tm),l,na)) :: genrec depth rest | Alias (initial,(na,c,d)) :: rest -> (* [c] is out of relocation scope *) - Alias (initial,(na,c,map_pair (relocate_index n1 n2 depth) d)) :: genrec depth rest + Alias (initial,(na,c,map_pair (relocate_index sigma n1 n2 depth) d)) :: genrec depth rest | NonDepAlias :: rest -> NonDepAlias :: genrec depth rest | Abstract (i,d) :: rest -> let i = relocate_rel n1 n2 depth i in - Abstract (i, RelDecl.map_constr (relocate_index n1 n2 depth) d) + Abstract (i, RelDecl.map_constr (fun c -> EConstr.Unsafe.to_constr (relocate_index sigma n1 n2 depth (EConstr.of_constr c))) d) :: genrec (depth+1) rest in genrec 0 (* [replace_tomatch n c tomatch] replaces [Rel n] by [c] in [tomatch] *) -let rec replace_term n c k t = - if isRel t && Int.equal (destRel t) (n + k) then lift k c - else map_constr_with_binders succ (replace_term n c) k t +let rec replace_term sigma n c k t = + if isRel sigma t && Int.equal (destRel sigma t) (n + k) then Vars.lift k c + else EConstr.map_with_binders sigma succ (replace_term sigma n c) k t let length_of_tomatch_type_sign na t = let l = match na with @@ -628,21 +648,21 @@ let length_of_tomatch_type_sign na t = | NotInd _ -> l | IsInd (_, _, names) -> List.length names + l -let replace_tomatch n c = +let replace_tomatch sigma n c = let rec replrec depth = function | [] -> [] | Pushed (initial,((b,tm),l,na)) :: rest -> - let b = replace_term n c depth b in - let tm = map_tomatch_type (replace_term n c depth) tm in + let b = replace_term sigma n c depth b in + let tm = map_tomatch_type (replace_term sigma n c depth) tm in List.iter (fun i -> if Int.equal i (n + depth) then anomaly (Pp.str "replace_tomatch")) l; Pushed (initial,((b,tm),l,na)) :: replrec depth rest | Alias (initial,(na,b,d)) :: rest -> (* [b] is out of replacement scope *) - Alias (initial,(na,b,map_pair (replace_term n c depth) d)) :: replrec depth rest + Alias (initial,(na,b,map_pair (replace_term sigma n c depth) d)) :: replrec depth rest | NonDepAlias :: rest -> NonDepAlias :: replrec depth rest | Abstract (i,d) :: rest -> - Abstract (i, RelDecl.map_constr (replace_term n c depth) d) + Abstract (i, RelDecl.map_constr (fun t -> EConstr.Unsafe.to_constr (replace_term sigma n c depth (EConstr.of_constr t))) d) :: replrec (depth+1) rest in replrec 0 @@ -667,7 +687,7 @@ let rec liftn_tomatch_stack n depth = function NonDepAlias :: liftn_tomatch_stack n depth rest | Abstract (i,d)::rest -> let i = if i map_predicate f (k+1) ccl rest -let noccur_predicate_between n = map_predicate (noccur_between n) +let noccur_predicate_between sigma n = map_predicate (noccur_between sigma n) let liftn_predicate n = map_predicate (liftn n) let lift_predicate n = liftn_predicate n 1 -let regeneralize_index_predicate n = map_predicate (relocate_index n 1) 0 +let regeneralize_index_predicate sigma n = map_predicate (relocate_index sigma n 1) 0 let substnl_predicate sigma = map_predicate (substnl sigma) @@ -857,7 +877,7 @@ let specialize_predicate_var (cur,typ,dep) tms ccl = let l = match typ with | IsInd (_, IndType (_, _), []) -> [] - | IsInd (_, IndType (_, realargs), names) -> realargs + | IsInd (_, IndType (_, realargs), names) -> List.map EConstr.of_constr realargs | NotInd _ -> [] in subst_predicate (l,c) ccl tms @@ -870,13 +890,13 @@ let specialize_predicate_var (cur,typ,dep) tms ccl = (* We first need to lift t(x) s.t. it is typed in Gamma, X:=rargs, x' *) (* then we have to replace x by x' in t(x) and y by y' in P *) (*****************************************************************************) -let generalize_predicate (names,na) ny d tms ccl = +let generalize_predicate sigma (names,na) ny d tms ccl = let () = match na with | Anonymous -> anomaly (Pp.str "Undetected dependency") | _ -> () in let p = List.length names + 1 in let ccl = lift_predicate 1 ccl tms in - regeneralize_index_predicate (ny+p+1) ccl tms + regeneralize_index_predicate sigma (ny+p+1) ccl tms (*****************************************************************************) (* We just matched over cur:ind(realargs) in the following matching problem *) @@ -906,7 +926,7 @@ let rec extract_predicate ccl = function subst1 cur pred end | Pushed (_,((cur,IsInd (_,IndType(_,realargs),_)),_,na))::tms -> - let realargs = List.rev realargs in + let realargs = List.rev_map EConstr.of_constr realargs in let k, nrealargs = match na with | Anonymous -> 0, realargs | Name _ -> 1, (cur :: realargs) @@ -925,9 +945,9 @@ let abstract_predicate env sigma indf cur realargs (names,na) tms ccl = (* that are rels, consistently with the specialization made in *) (* build_branch *) let tms = List.fold_right2 (fun par arg tomatch -> - match kind_of_term par with - | Rel i -> relocate_index_tomatch (i+n) (destRel arg) tomatch - | _ -> tomatch) (realargs@[cur]) (Context.Rel.to_extended_list 0 sign) + match EConstr.kind sigma par with + | Rel i -> relocate_index_tomatch sigma (i+n) (destRel sigma arg) tomatch + | _ -> tomatch) (realargs@[cur]) (List.map EConstr.of_constr (Context.Rel.to_extended_list 0 sign)) (lift_tomatch_stack n tms) in (* Pred is already dependent in the current term to match (if *) (* (na<>Anonymous) and its realargs; we just need to adjust it to *) @@ -939,7 +959,7 @@ let abstract_predicate env sigma indf cur realargs (names,na) tms ccl = let pred = extract_predicate ccl tms in (* Build the predicate properly speaking *) let sign = List.map2 set_name (na::names) sign in - it_mkLambda_or_LetIn_name env pred sign + EConstr.of_constr (it_mkLambda_or_LetIn_name env (EConstr.Unsafe.to_constr pred) sign) (* [expand_arg] is used by [specialize_predicate] if Yk denotes [Xk;xk] or [Xk], @@ -974,6 +994,10 @@ let add_assert_false_case pb tomatch = let adjust_impossible_cases pb pred tomatch submat = match submat with | [] -> + (** FIXME: This breaks if using evar-insensitive primitives. In particular, + this means that the Evd.define below may redefine an already defined + evar. See e.g. first definition of test for bug #3388. *) + let pred = EConstr.Unsafe.to_constr pred in begin match kind_of_term pred with | Evar (evk,_) when snd (evar_source evk !(pb.evdref)) == Evar_kinds.ImpossibleCase -> if not (Evd.is_defined !(pb.evdref) evk) then begin @@ -1024,27 +1048,30 @@ let specialize_predicate newtomatchs (names,depna) arsign cs tms ccl = (* We prepare the substitution of X and x:I(X) *) let realargsi = if not (Int.equal nrealargs 0) then - subst_of_rel_context_instance arsign (Array.to_list cs.cs_concl_realargs) + CVars.subst_of_rel_context_instance arsign (Array.to_list cs.cs_concl_realargs) else [] in + let realargsi = List.map EConstr.of_constr realargsi in let copti = match depna with | Anonymous -> None - | Name _ -> Some (build_dependent_constructor cs) + | Name _ -> Some (EConstr.of_constr (build_dependent_constructor cs)) in (* The substituends realargsi, copti are all defined in gamma, x1...xn *) (* We need _parallel_ bindings to get gamma, x1...xn |- PI tms. ccl'' *) (* Note: applying the substitution in tms is not important (is it sure?) *) let ccl'' = - whd_betaiota Evd.empty (EConstr.of_constr (subst_predicate (realargsi, copti) ccl' tms)) in + whd_betaiota Evd.empty (subst_predicate (realargsi, copti) ccl' tms) in + let ccl'' = EConstr.of_constr ccl'' in (* We adjust ccl st: gamma, x'1..x'n, x1..xn, tms |- ccl'' *) let ccl''' = liftn_predicate n (n+1) ccl'' tms in (* We finally get gamma,x'1..x'n,x |- [X1;x1:I(X1)]..[Xn;xn:I(Xn)]pred'''*) snd (List.fold_left (expand_arg tms) (1,ccl''') newtomatchs) let find_predicate loc env evdref p current (IndType (indf,realargs)) dep tms = + let realargs = List.map EConstr.of_constr realargs in let pred = abstract_predicate env !evdref indf current realargs dep tms p in - (pred, whd_betaiota !evdref - (EConstr.of_constr (applist (pred, realargs@[current])))) + (pred, EConstr.of_constr (whd_betaiota !evdref + (applist (pred, realargs@[current])))) (* Take into account that a type has been discovered to be inductive, leading to more dependencies in the predicate if the type has indices *) @@ -1065,40 +1092,40 @@ let adjust_predicate_from_tomatch tomatch (current,typ as ct) pb = (* Remove commutative cuts that turn out to be non-dependent after some evars have been instantiated *) -let rec ungeneralize n ng body = - match kind_of_term body with +let rec ungeneralize sigma n ng body = + match EConstr.kind sigma body with | Lambda (_,_,c) when Int.equal ng 0 -> subst1 (mkRel n) c | Lambda (na,t,c) -> (* We traverse an inner generalization *) - mkLambda (na,t,ungeneralize (n+1) (ng-1) c) + mkLambda (na,t,ungeneralize sigma (n+1) (ng-1) c) | LetIn (na,b,t,c) -> (* We traverse an alias *) - mkLetIn (na,b,t,ungeneralize (n+1) ng c) + mkLetIn (na,b,t,ungeneralize sigma (n+1) ng c) | Case (ci,p,c,brs) -> (* We traverse a split *) let p = - let sign,p = decompose_lam_assum p in - let sign2,p = decompose_prod_n_assum ng p in - let p = prod_applist p [mkRel (n+List.length sign+ng)] in + let sign,p = decompose_lam_assum sigma p in + let sign2,p = decompose_prod_n_assum sigma ng p in + let p = prod_applist sigma p [mkRel (n+List.length sign+ng)] in it_mkLambda_or_LetIn (it_mkProd_or_LetIn p sign2) sign in mkCase (ci,p,c,Array.map2 (fun q c -> - let sign,b = decompose_lam_n_decls q c in - it_mkLambda_or_LetIn (ungeneralize (n+q) ng b) sign) + let sign,b = decompose_lam_n_decls sigma q c in + it_mkLambda_or_LetIn (ungeneralize sigma (n+q) ng b) sign) ci.ci_cstr_ndecls brs) | App (f,args) -> (* We traverse an inner generalization *) - assert (isCase f); - mkApp (ungeneralize n (ng+Array.length args) f,args) + assert (isCase sigma f); + mkApp (ungeneralize sigma n (ng+Array.length args) f,args) | _ -> assert false -let ungeneralize_branch n k (sign,body) cs = - (sign,ungeneralize (n+cs.cs_nargs) k body) +let ungeneralize_branch sigma n k (sign,body) cs = + (sign,ungeneralize sigma (n+cs.cs_nargs) k body) let rec is_dependent_generalization sigma ng body = - match kind_of_term body with + match EConstr.kind sigma body with | Lambda (_,_,c) when Int.equal ng 0 -> - not (EConstr.Vars.noccurn sigma 1 (EConstr.of_constr c)) + not (noccurn sigma 1 c) | Lambda (na,t,c) -> (* We traverse an inner generalization *) is_dependent_generalization sigma (ng-1) c @@ -1108,12 +1135,12 @@ let rec is_dependent_generalization sigma ng body = | Case (ci,p,c,brs) -> (* We traverse a split *) Array.exists2 (fun q c -> - let _,b = decompose_lam_n_decls q c in + let _,b = decompose_lam_n_decls sigma q c in is_dependent_generalization sigma ng b) ci.ci_cstr_ndecls brs | App (g,args) -> (* We traverse an inner generalization *) - assert (isCase g); + assert (isCase sigma g); is_dependent_generalization sigma (ng+Array.length args) g | _ -> assert false @@ -1140,9 +1167,9 @@ let postprocess_dependencies evd tocheck brs tomatch pred deps cs = (* terms by its actual value in both the remaining terms to match and *) (* the bodies of the Case *) let pred = lift_predicate (-1) pred tomatch in - let tomatch = relocate_index_tomatch 1 (n+1) tomatch in + let tomatch = relocate_index_tomatch evd 1 (n+1) tomatch in let tomatch = lift_tomatch_stack (-1) tomatch in - let brs = Array.map2 (ungeneralize_branch n k) brs cs in + let brs = Array.map2 (ungeneralize_branch evd n k) brs cs in aux k brs tomatch pred tocheck deps | _ -> assert false in aux 0 brs tomatch pred tocheck deps @@ -1194,24 +1221,24 @@ let rec generalize_problem names pb = function | [] -> pb, [] | i::l -> let pb',deps = generalize_problem names pb l in - let d = map_constr (lift i) (Environ.lookup_rel i pb.env) in + let d = map_constr (CVars.lift i) (Environ.lookup_rel i pb.env) in begin match d with | LocalDef (Anonymous,_,_) -> pb', deps | _ -> (* for better rendering *) let d = RelDecl.map_type (fun c -> whd_betaiota !(pb.evdref) (EConstr.of_constr c)) d in let tomatch = lift_tomatch_stack 1 pb'.tomatch in - let tomatch = relocate_index_tomatch (i+1) 1 tomatch in + let tomatch = relocate_index_tomatch !(pb.evdref) (i+1) 1 tomatch in { pb' with tomatch = Abstract (i,d) :: tomatch; - pred = generalize_predicate names i d pb'.tomatch pb'.pred }, + pred = generalize_predicate !(pb'.evdref) names i d pb'.tomatch pb'.pred }, i::deps end (* No more patterns: typing the right-hand side of equations *) let build_leaf pb = let rhs = extract_rhs pb in - let j = pb.typing_function (mk_tycon (EConstr.of_constr pb.pred)) rhs.rhs_env pb.evdref rhs.it in + let j = pb.typing_function (mk_tycon pb.pred) rhs.rhs_env pb.evdref rhs.it in j_nf_evar !(pb.evdref) j (* Build the sub-pattern-matching problem for a given branch "C x1..xn as x" *) @@ -1238,7 +1265,7 @@ let build_branch initial current realargs deps (realnames,curname) pb arsign eqn (* We adjust the terms to match in the context they will be once the *) (* context [x1:T1,..,xn:Tn] will have been pushed on the current env *) let typs' = - List.map_i (fun i d -> (mkRel i, map_constr (lift i) d)) 1 typs in + List.map_i (fun i d -> (mkRel i, map_constr (CVars.lift i) d)) 1 typs in let extenv = push_rel_context typs pb.env in @@ -1255,24 +1282,24 @@ let build_branch initial current realargs deps (realnames,curname) pb arsign eqn (* The dependent term to subst in the types of the remaining UnPushed terms is relative to the current context enriched by topushs *) - let ci = build_dependent_constructor const_info in + let ci = EConstr.of_constr (build_dependent_constructor const_info) in (* Current context Gamma has the form Gamma1;cur:I(realargs);Gamma2 *) (* We go from Gamma |- PI tms. pred to *) (* Gamma;x1..xn;curalias:I(x1..xn) |- PI tms'. pred' *) (* where, in tms and pred, those realargs that are vars are *) (* replaced by the corresponding xi and cur replaced by curalias *) - let cirealargs = Array.to_list const_info.cs_concl_realargs in + let cirealargs = Array.map_to_list EConstr.of_constr const_info.cs_concl_realargs in (* Do the specialization for terms to match *) let tomatch = List.fold_right2 (fun par arg tomatch -> - match kind_of_term par with - | Rel i -> replace_tomatch (i+const_info.cs_nargs) arg tomatch + match EConstr.kind !(pb.evdref) par with + | Rel i -> replace_tomatch !(pb.evdref) (i+const_info.cs_nargs) arg tomatch | _ -> tomatch) (current::realargs) (ci::cirealargs) (lift_tomatch_stack const_info.cs_nargs pb.tomatch) in let pred_is_not_dep = - noccur_predicate_between 1 (List.length realnames + 1) pb.pred tomatch in + noccur_predicate_between !(pb.evdref) 1 (List.length realnames + 1) pb.pred tomatch in let typs' = List.map2 @@ -1298,10 +1325,10 @@ let build_branch initial current realargs deps (realnames,curname) pb arsign eqn | Name _ -> let cur_alias = lift const_info.cs_nargs current in let ind = - appvect ( + mkApp ( applist (mkIndU (inductive_of_constructor (fst const_info.cs_cstr), snd const_info.cs_cstr), - List.map (lift const_info.cs_nargs) const_info.cs_params), - const_info.cs_concl_realargs) in + List.map (EConstr.of_constr %> lift const_info.cs_nargs) const_info.cs_params), + Array.map EConstr.of_constr const_info.cs_concl_realargs) in Alias (initial,(aliasname,cur_alias,(ci,ind))) in let tomatch = List.rev_append (alias :: currents) tomatch in @@ -1361,13 +1388,14 @@ and match_current pb (initial,tomatch) = if (not no_cstr || not (List.is_empty pb.mat)) && onlydflt then compile_all_variables initial tomatch pb else + let realargs = List.map EConstr.of_constr realargs in (* We generalize over terms depending on current term to match *) let pb,deps = generalize_problem (names,dep) pb deps in (* We compile branches *) let brvals = Array.map2 (compile_branch initial current realargs (names,dep) deps pb arsign) eqns cstrs in (* We build the (elementary) case analysis *) - let depstocheck = current::binding_vars_of_inductive typ in + let depstocheck = current::binding_vars_of_inductive !(pb.evdref) typ in let brvals,tomatch,pred,inst = postprocess_dependencies !(pb.evdref) depstocheck brvals pb.tomatch pb.pred deps cstrs in @@ -1377,13 +1405,14 @@ and match_current pb (initial,tomatch) = find_predicate pb.caseloc pb.env pb.evdref pred current indt (names,dep) tomatch in let ci = make_case_info pb.env (fst mind) pb.casestyle in - let pred = nf_betaiota !(pb.evdref) (EConstr.of_constr pred) in + let pred = nf_betaiota !(pb.evdref) pred in + let pred = EConstr.of_constr pred in let case = - make_case_or_project pb.env indf ci pred current brvals + make_case_or_project pb.env !(pb.evdref) indf ci pred current brvals in - Typing.check_allowed_sort pb.env !(pb.evdref) mind (EConstr.of_constr current) (EConstr.of_constr pred); - { uj_val = applist (case, inst); - uj_type = prod_applist typ inst } + Typing.check_allowed_sort pb.env !(pb.evdref) mind current pred; + { uj_val = EConstr.Unsafe.to_constr (applist (case, inst)); + uj_type = EConstr.Unsafe.to_constr (prod_applist !(pb.evdref) typ inst) } (* Building the sub-problem when all patterns are variables. Case @@ -1394,14 +1423,14 @@ and shift_problem ((current,t),_,na) pb = let pred = specialize_predicate_var (current,t,na) pb.tomatch pb.pred in let pb = { pb with - env = push_rel (LocalDef (na,current,ty)) pb.env; + env = push_rel (local_def (na,current,ty)) pb.env; tomatch = tomatch; pred = lift_predicate 1 pred tomatch; history = pop_history pb.history; mat = List.map (push_current_pattern (current,ty)) pb.mat } in let j = compile pb in - { uj_val = subst1 current j.uj_val; - uj_type = subst1 current j.uj_type } + { uj_val = EConstr.Unsafe.to_constr (subst1 current (EConstr.of_constr j.uj_val)); + uj_type = EConstr.Unsafe.to_constr (subst1 current (EConstr.of_constr j.uj_type)) } (* Building the sub-problem when all patterns are variables, non-initial case. Variables which appear as subterms of constructor @@ -1424,7 +1453,7 @@ and compile_all_variables initial cur pb = (* Building the sub-problem when all patterns are variables *) and compile_branch initial current realargs names deps pb arsign eqns cstr = let sign, pb = build_branch initial current realargs deps names pb arsign eqns cstr in - sign, (compile pb).uj_val + sign, EConstr.of_constr (compile pb).uj_val (* Abstract over a declaration before continuing splitting *) and compile_generalization pb i d rest = @@ -1434,15 +1463,15 @@ and compile_generalization pb i d rest = tomatch = rest; mat = List.map (push_generalized_decl_eqn pb.env i d) pb.mat } in let j = compile pb in - { uj_val = mkLambda_or_LetIn d j.uj_val; - uj_type = mkProd_wo_LetIn d j.uj_type } + { uj_val = Term.mkLambda_or_LetIn d j.uj_val; + uj_type = Term.mkProd_wo_LetIn d j.uj_type } (* spiwack: the [initial] argument keeps track whether the alias has been introduced by a toplevel branch ([true]) or a deep one ([false]). *) and compile_alias initial pb (na,orig,(expanded,expanded_typ)) rest = let f c t = - let alias = LocalDef (na,c,t) in + let alias = local_def (na,c,t) in let pb = { pb with env = push_rel alias pb.env; @@ -1451,12 +1480,13 @@ and compile_alias initial pb (na,orig,(expanded,expanded_typ)) rest = history = pop_history_pattern pb.history; mat = List.map (push_alias_eqn alias) pb.mat } in let j = compile pb in + let sigma = !(pb.evdref) in { uj_val = - if isRel c || isVar c || count_occurrences !(pb.evdref) (EConstr.mkRel 1) (EConstr.of_constr j.uj_val) <= 1 then - subst1 c j.uj_val + if isRel sigma c || isVar sigma c || count_occurrences sigma (mkRel 1) (EConstr.of_constr j.uj_val) <= 1 then + EConstr.Unsafe.to_constr (subst1 c (EConstr.of_constr j.uj_val)) else - mkLetIn (na,c,t,j.uj_val); - uj_type = subst1 c j.uj_type } in + EConstr.Unsafe.to_constr (mkLetIn (na,c,t,EConstr.of_constr j.uj_val)); + uj_type = EConstr.Unsafe.to_constr (subst1 c (EConstr.of_constr j.uj_type)) } in (* spiwack: when an alias appears on a deep branch, its non-expanded form is automatically a variable of the same name. We avoid introducing such superfluous aliases so that refines are elegant. *) @@ -1477,10 +1507,10 @@ and compile_alias initial pb (na,orig,(expanded,expanded_typ)) rest = evaluation; the drawback is that it might duplicate the instances of the term to match when the corresponding variable is substituted by a non-evaluated expression *) - if not (Flags.is_program_mode ()) && (isRel orig || isVar orig) then + if not (Flags.is_program_mode ()) && (isRel sigma orig || isVar sigma orig) then (* Try to compile first using non expanded alias *) try - if initial then f orig (Retyping.get_type_of pb.env !(pb.evdref) (EConstr.of_constr orig)) + if initial then f orig (EConstr.of_constr (Retyping.get_type_of pb.env sigma orig)) else just_pop () with e when precatchable_exception e -> (* Try then to compile using expanded alias *) @@ -1495,7 +1525,7 @@ and compile_alias initial pb (na,orig,(expanded,expanded_typ)) rest = (* Could be needed in case of a recursive call which requires to be on a variable for size reasons *) pb.evdref := sigma; - if initial then f orig (Retyping.get_type_of pb.env !(pb.evdref) (EConstr.of_constr orig)) + if initial then f orig (EConstr.of_constr (Retyping.get_type_of pb.env !(pb.evdref) orig)) else just_pop () @@ -1579,11 +1609,11 @@ let adjust_to_extended_env_and_remove_deps env extenv sigma subst t = let (p, _, _) = lookup_rel_id x (rel_context extenv) in let rec traverse_local_defs p = match lookup_rel p extenv with - | LocalDef (_,c,_) -> assert (isRel c); traverse_local_defs (p + destRel c) + | LocalDef (_,c,_) -> assert (isRel sigma (EConstr.of_constr c)); traverse_local_defs (p + destRel sigma (EConstr.of_constr c)) | LocalAssum _ -> p in let p = traverse_local_defs p in let u = lift (n' - n) u in - try Some (p, u, EConstr.Unsafe.to_constr (expand_vars_in_term extenv sigma (EConstr.of_constr u))) + try Some (p, u, expand_vars_in_term extenv sigma u) (* pedrot: does this really happen to raise [Failure _]? *) with Failure _ -> None in let subst0 = List.map_filter map subst in @@ -1613,8 +1643,9 @@ let rec list_assoc_in_triple x = function *) let abstract_tycon loc env evdref subst tycon extenv t = - let t = nf_betaiota !evdref (EConstr.of_constr t) in (* it helps in some cases to remove K-redex*) - let src = match kind_of_term t with + let t = nf_betaiota !evdref t in (* it helps in some cases to remove K-redex*) + let t = EConstr.of_constr t in + let src = match EConstr.kind !evdref t with | Evar (evk,_) -> (loc,Evar_kinds.SubEvar evk) | _ -> (loc,Evar_kinds.CasesType true) in let subst0,t0 = adjust_to_extended_env_and_remove_deps env extenv !evdref subst t in @@ -1624,10 +1655,10 @@ let abstract_tycon loc env evdref subst tycon extenv t = by an evar that may depend (and only depend) on the corresponding convertible subterms of the substitution *) let rec aux (k,env,subst as x) t = - let t = whd_evar !evdref t in match kind_of_term t with + match EConstr.kind !evdref t with | Rel n when is_local_def (lookup_rel n env) -> t | Evar ev -> - let ty = get_type_of env !evdref (EConstr.of_constr t) in + let ty = get_type_of env !evdref t in let ty = Evarutil.evd_comb1 (refresh_universes (Some false) env) evdref (EConstr.of_constr ty) in let inst = List.map_i @@ -1635,39 +1666,43 @@ let abstract_tycon loc env evdref subst tycon extenv t = try list_assoc_in_triple i subst0 with Not_found -> mkRel i) 1 (rel_context env) in let ev' = e_new_evar env evdref ~src ty in - let ev = (fst ev, Array.map EConstr.of_constr (snd ev)) in - begin match solve_simple_eqn (evar_conv_x full_transparent_state) env !evdref (None,ev,EConstr.of_constr (substl inst ev')) with + let ev' = EConstr.of_constr ev' in + begin match solve_simple_eqn (evar_conv_x full_transparent_state) env !evdref (None,ev,substl inst ev') with | Success evd -> evdref := evd | UnifFailure _ -> assert false end; ev' | _ -> - let good = List.filter (fun (_,u,_) -> is_conv_leq env !evdref (EConstr.of_constr t) (EConstr.of_constr u)) subst in + let good = List.filter (fun (_,u,_) -> is_conv_leq env !evdref t u) subst in match good with | [] -> - let self env c = EConstr.of_constr (aux env (EConstr.Unsafe.to_constr c)) in - EConstr.Unsafe.to_constr (map_constr_with_full_binders !evdref push_binder self x (EConstr.of_constr t)) + map_constr_with_full_binders !evdref push_binder aux x t | (_, _, u) :: _ -> (* u is in extenv *) let vl = List.map pi1 good in let ty = - let ty = get_type_of env !evdref (EConstr.of_constr t) in - Evarutil.evd_comb1 (refresh_universes (Some false) env) evdref (EConstr.of_constr ty) + let ty = get_type_of env !evdref t in + let ty = EConstr.of_constr ty in + Evarutil.evd_comb1 (refresh_universes (Some false) env) evdref ty in + let ty = EConstr.of_constr ty in let ty = lift (-k) (aux x ty) in - let depvl = free_rels !evdref (EConstr.of_constr ty) in + let depvl = free_rels !evdref ty in let inst = List.map_i (fun i _ -> if Int.List.mem i vl then u else mkRel i) 1 (rel_context extenv) in let rel_filter = - List.map (fun a -> not (isRel a) || dependent !evdref (EConstr.of_constr a) (EConstr.of_constr u) - || Int.Set.mem (destRel a) depvl) inst in + List.map (fun a -> not (isRel !evdref a) || dependent !evdref a u + || Int.Set.mem (destRel !evdref a) depvl) inst in let named_filter = - List.map (fun d -> local_occur_var !evdref (NamedDecl.get_id d) (EConstr.of_constr u)) + List.map (fun d -> local_occur_var !evdref (NamedDecl.get_id d) u) (named_context extenv) in let filter = Filter.make (rel_filter @ named_filter) in let candidates = u :: List.map mkRel vl in + let candidates = List.map EConstr.Unsafe.to_constr candidates in + let ty = EConstr.Unsafe.to_constr ty in let ev = e_new_evar extenv evdref ~src ~filter ~candidates ty in + let ev = EConstr.of_constr ev in lift k ev in aux (0,extenv,subst0) t0 @@ -1681,15 +1716,17 @@ let build_tycon loc env tycon_env s subst tycon extenv evdref t = let n' = Context.Rel.length (rel_context tycon_env) in let impossible_case_type, u = e_new_type_evar (reset_context env) evdref univ_flexible_alg ~src:(loc,Evar_kinds.ImpossibleCase) in + let impossible_case_type = EConstr.of_constr impossible_case_type in (lift (n'-n) impossible_case_type, mkSort u) | Some t -> let t = abstract_tycon loc tycon_env evdref subst tycon extenv t in - let evd,tt = Typing.type_of extenv !evdref (EConstr.of_constr t) in + let evd,tt = Typing.type_of extenv !evdref t in + let tt = EConstr.of_constr tt in evdref := evd; (t,tt) in - let b = e_cumul env evdref (EConstr.of_constr tt) (EConstr.mkSort s) (* side effect *) in + let b = e_cumul env evdref tt (mkSort s) (* side effect *) in if not b then anomaly (Pp.str "Build_tycon: should be a type"); - { uj_val = t; uj_type = tt } + { uj_val = EConstr.Unsafe.to_constr t; uj_type = EConstr.Unsafe.to_constr tt } (* For a multiple pattern-matching problem Xi on t1..tn with return * type T, [build_inversion_problem Gamma Sigma (t1..tn) T] builds a return @@ -1703,13 +1740,13 @@ let build_tycon loc env tycon_env s subst tycon extenv evdref t = let build_inversion_problem loc env sigma tms t = let make_patvar t (subst,avoid) = - let id = next_name_away (named_hd env t Anonymous) avoid in + let id = next_name_away (named_hd env (EConstr.Unsafe.to_constr t) Anonymous) avoid in PatVar (Loc.ghost,Name id), ((id,t)::subst, id::avoid) in let rec reveal_pattern t (subst,avoid as acc) = - match kind_of_term (whd_all env sigma (EConstr.of_constr t)) with + match EConstr.kind sigma (EConstr.of_constr (whd_all env sigma t)) with | Construct (cstr,u) -> PatCstr (Loc.ghost,cstr,[],Anonymous), acc - | App (f,v) when isConstruct f -> - let cstr,u = destConstruct f in + | App (f,v) when isConstruct sigma f -> + let cstr,u = destConstruct sigma f in let n = constructor_nrealargs_env env cstr in let l = List.lastn n (Array.to_list v) in let l,acc = List.fold_map' reveal_pattern l acc in @@ -1719,6 +1756,7 @@ let build_inversion_problem loc env sigma tms t = match tms with | [] -> [], acc_sign, acc | (t, IsInd (_,IndType(indf,realargs),_)) :: tms -> + let realargs = List.map EConstr.of_constr realargs in let patl,acc = List.fold_map' reveal_pattern realargs acc in let pat,acc = make_patvar t acc in let indf' = lift_inductive_family n indf in @@ -1731,7 +1769,7 @@ let build_inversion_problem loc env sigma tms t = List.rev_append patl patl',acc_sign,acc | (t, NotInd (bo,typ)) :: tms -> let pat,acc = make_patvar t acc in - let d = LocalAssum (alias_of_pat pat,typ) in + let d = local_assum (alias_of_pat pat,typ) in let patl,acc_sign,acc = aux (n+1) (push_rel d env) (d::acc_sign) tms acc in pat::patl,acc_sign,acc in let avoid0 = ids_of_context env in @@ -1748,7 +1786,7 @@ let build_inversion_problem loc env sigma tms t = let n = List.length sign in let decls = - List.map_i (fun i d -> (mkRel i, map_constr (lift i) d)) 1 sign in + List.map_i (fun i d -> (mkRel i, map_constr (CVars.lift i) d)) 1 sign in let pb_env = push_rel_context sign env in let decls = @@ -1799,7 +1837,7 @@ let build_inversion_problem loc env sigma tms t = it = None } } ] in (* [pb] is the auxiliary pattern-matching serving as skeleton for the return type of the original problem Xi *) - let s' = Retyping.get_sort_of env sigma (EConstr.of_constr t) in + let s' = Retyping.get_sort_of env sigma t in let sigma, s = Evd.new_sort_variable univ_flexible_alg sigma in let sigma = Evd.set_leq_sort env sigma s' s in let evdref = ref sigma in @@ -1813,7 +1851,7 @@ let build_inversion_problem loc env sigma tms t = caseloc = loc; casestyle = RegularStyle; typing_function = build_tycon loc env pb_env s subst} in - let pred = (compile pb).uj_val in + let pred = EConstr.of_constr (compile pb).uj_val in (!evdref,pred) (* Here, [pred] is assumed to be in the context built from all *) @@ -1835,8 +1873,8 @@ let extract_arity_signature ?(dolift=true) env0 tomatchl tmsign = | NotInd (bo,typ) -> (match t with | None -> (match bo with - | None -> [LocalAssum (na, lift n typ)] - | Some b -> [LocalDef (na, lift n b, lift n typ)]) + | None -> [local_assum (na, lift n typ)] + | Some b -> [local_def (na, lift n b, lift n typ)]) | Some (loc,_,_) -> user_err ~loc (str"Unexpected type annotation for a term of non inductive type.")) @@ -1879,8 +1917,8 @@ let prepare_predicate_from_arsign_tycon env sigma loc tomatchs arsign c = let subst, len = List.fold_right2 (fun (tm, tmtype) sign (subst, len) -> let signlen = List.length sign in - match kind_of_term tm with - | Rel n when dependent sigma (EConstr.of_constr tm) (EConstr.of_constr c) + match EConstr.kind sigma tm with + | Rel n when dependent sigma tm c && Int.equal signlen 1 (* The term to match is not of a dependent type itself *) -> ((n, len) :: subst, len - signlen) | Rel n when signlen > 1 (* The term is of a dependent type, @@ -1888,24 +1926,25 @@ let prepare_predicate_from_arsign_tycon env sigma loc tomatchs arsign c = (match tmtype with NotInd _ -> (subst, len - signlen) | IsInd (_, IndType(indf,realargs),_) -> + let realargs = List.map EConstr.of_constr realargs in let subst, len = List.fold_left (fun (subst, len) arg -> - match kind_of_term arg with - | Rel n when dependent sigma (EConstr.of_constr arg) (EConstr.of_constr c) -> + match EConstr.kind sigma arg with + | Rel n when dependent sigma arg c -> ((n, len) :: subst, pred len) | _ -> (subst, pred len)) (subst, len) realargs in let subst = - if dependent sigma (EConstr.of_constr tm) (EConstr.of_constr c) && List.for_all isRel realargs + if dependent sigma tm c && List.for_all (isRel sigma) realargs then (n, len) :: subst else subst in (subst, pred len)) | _ -> (subst, len - signlen)) (List.rev tomatchs) arsign ([], nar) in let rec predicate lift c = - match kind_of_term c with + match EConstr.kind sigma c with | Rel n when n > lift -> (try (* Make the predicate dependent on the matched variable *) @@ -1915,12 +1954,12 @@ let prepare_predicate_from_arsign_tycon env sigma loc tomatchs arsign c = (* A variable that is not matched, lift over the arsign. *) mkRel (n + nar)) | _ -> - map_constr_with_binders succ predicate lift c + EConstr.map_with_binders sigma succ predicate lift c in assert (len == 0); let p = predicate 0 c in let env' = List.fold_right push_rel_context arsign env in - try let sigma' = fst (Typing.type_of env' sigma (EConstr.of_constr p)) in + try let sigma' = fst (Typing.type_of env' sigma p) in Some (sigma', p) with e when precatchable_exception e -> None @@ -1935,11 +1974,26 @@ let prepare_predicate_from_arsign_tycon env sigma loc tomatchs arsign c = * tycon to make the predicate if it is not closed. *) +exception LocalOccur + +let noccur_with_meta sigma n m term = + let rec occur_rec n c = match EConstr.kind sigma c with + | Rel p -> if n<=p && p + (match EConstr.kind sigma f with + | Cast (c,_,_) when isMeta sigma c -> () + | Meta _ -> () + | _ -> EConstr.iter_with_binders sigma succ occur_rec n c) + | Evar (_, _) -> () + | _ -> EConstr.iter_with_binders sigma succ occur_rec n c + in + try (occur_rec n term; true) with LocalOccur -> false + let prepare_predicate loc typing_fun env sigma tomatchs arsign tycon pred = let preds = match pred, tycon with (* No return clause *) - | None, Some t when not (noccur_with_meta 0 max_int t) -> + | None, Some t when not (noccur_with_meta sigma 0 max_int t) -> (* If the tycon is not closed w.r.t real variables, we try *) (* two different strategies *) (* First strategy: we abstract the tycon wrt to the dependencies *) @@ -1960,7 +2014,7 @@ let prepare_predicate loc typing_fun env sigma tomatchs arsign tycon pred = let Sigma ((t, _), sigma, _) = new_type_evar env sigma univ_flexible_alg ~src:(loc, Evar_kinds.CasesType false) in let sigma = Sigma.to_evar_map sigma in - sigma, t + sigma, EConstr.of_constr t in (* First strategy: we build an "inversion" predicate *) let sigma1,pred1 = build_inversion_problem loc env sigma tomatchs t in @@ -1975,7 +2029,7 @@ let prepare_predicate loc typing_fun env sigma tomatchs arsign tycon pred = let evdref = ref sigma in let predcclj = typing_fun (mk_tycon (EConstr.mkSort newt)) envar evdref rtntyp in let sigma = !evdref in - let predccl = (j_nf_evar sigma predcclj).uj_val in + let predccl = EConstr.of_constr (j_nf_evar sigma predcclj).uj_val in [sigma, predccl] in List.map @@ -2013,7 +2067,6 @@ let eq_id avoid id = let hid' = next_ident_away hid avoid in hid' -let papp evdref gr args = EConstr.Unsafe.to_constr (papp evdref gr (Array.map EConstr.of_constr args)) let mk_eq evdref typ x y = papp evdref coq_eq_ind [| typ; x ; y |] let mk_eq_refl evdref typ x = papp evdref coq_eq_refl [| typ; x |] let mk_JMeq evdref typ x typ' y = @@ -2035,16 +2088,17 @@ let constr_of_pat env evdref arsign pat avoid = let previd, id = prime avoid (Name (Id.of_string "wildcard")) in Name id, id :: avoid in - (PatVar (l, name), [LocalAssum (name, ty)] @ realargs, mkRel 1, ty, + (PatVar (l, name), [local_assum (name, ty)] @ realargs, mkRel 1, ty, (List.map (fun x -> mkRel 1) realargs), 1, avoid) | PatCstr (l,((_, i) as cstr),args,alias) -> let cind = inductive_of_constructor cstr in let IndType (indf, _) = - try find_rectype env ( !evdref) (EConstr.of_constr (lift (-(List.length realargs)) ty)) + try find_rectype env ( !evdref) (lift (-(List.length realargs)) ty) with Not_found -> error_case_not_inductive env !evdref - {uj_val = ty; uj_type = Typing.unsafe_type_of env !evdref (EConstr.of_constr ty)} + {uj_val = EConstr.Unsafe.to_constr ty; uj_type = Typing.unsafe_type_of env !evdref ty} in let (ind,u), params = dest_ind_family indf in + let params = List.map EConstr.of_constr params in if not (eq_ind ind cind) then error_bad_constructor ~loc:l env cstr ind; let cstrs = get_constructors env indf in let ci = cstrs.(i-1) in @@ -2053,7 +2107,7 @@ let constr_of_pat env evdref arsign pat avoid = let patargs, args, sign, env, n, m, avoid = List.fold_right2 (fun decl ua (patargs, args, sign, env, n, m, avoid) -> - let t = RelDecl.get_type decl in + let t = EConstr.of_constr (RelDecl.get_type decl) in let pat', sign', arg', typ', argtypargs, n', avoid = let liftt = liftn (List.length sign) (succ (List.length args)) t in typ env (substl args liftt, []) ua avoid @@ -2067,34 +2121,36 @@ let constr_of_pat env evdref arsign pat avoid = let patargs = List.rev patargs in let pat' = PatCstr (l, cstr, patargs, alias) in let cstr = mkConstructU ci.cs_cstr in - let app = applistc cstr (List.map (lift (List.length sign)) params) in - let app = applistc app args in - let apptype = Retyping.get_type_of env ( !evdref) (EConstr.of_constr app) in - let IndType (indf, realargs) = find_rectype env (!evdref) (EConstr.of_constr apptype) in + let app = applist (cstr, List.map (lift (List.length sign)) params) in + let app = applist (app, args) in + let apptype = Retyping.get_type_of env ( !evdref) app in + let apptype = EConstr.of_constr apptype in + let IndType (indf, realargs) = find_rectype env (!evdref) apptype in + let realargs = List.map EConstr.of_constr realargs in match alias with Anonymous -> pat', sign, app, apptype, realargs, n, avoid | Name id -> - let sign = LocalAssum (alias, lift m ty) :: sign in + let sign = local_assum (alias, lift m ty) :: sign in let avoid = id :: avoid in let sign, i, avoid = try let env = push_rel_context sign env in evdref := the_conv_x_leq (push_rel_context sign env) - (EConstr.of_constr (lift (succ m) ty)) (EConstr.of_constr (lift 1 apptype)) !evdref; + (lift (succ m) ty) (lift 1 apptype) !evdref; let eq_t = mk_eq evdref (lift (succ m) ty) (mkRel 1) (* alias *) (lift 1 app) (* aliased term *) in let neq = eq_id avoid id in - LocalDef (Name neq, mkRel 0, eq_t) :: sign, 2, neq :: avoid + local_def (Name neq, mkRel 0, eq_t) :: sign, 2, neq :: avoid with Reduction.NotConvertible -> sign, 1, avoid in (* Mark the equality as a hole *) pat', sign, lift i app, lift i apptype, realargs, n + i, avoid in - let pat', sign, patc, patty, args, z, avoid = typ env (RelDecl.get_type (List.hd arsign), List.tl arsign) pat avoid in - pat', (sign, patc, (RelDecl.get_type (List.hd arsign), args), pat'), avoid + let pat', sign, patc, patty, args, z, avoid = typ env (EConstr.of_constr (RelDecl.get_type (List.hd arsign)), List.tl arsign) pat avoid in + pat', (sign, patc, (EConstr.of_constr (RelDecl.get_type (List.hd arsign)), args), pat'), avoid (* shadows functional version *) @@ -2104,22 +2160,22 @@ let eq_id avoid id = avoid := hid' :: !avoid; hid' -let is_topvar t = -match kind_of_term t with +let is_topvar sigma t = +match EConstr.kind sigma t with | Rel 0 -> true | _ -> false -let rels_of_patsign = +let rels_of_patsign sigma = List.map (fun decl -> match decl with - | LocalDef (na,t',t) when is_topvar t' -> LocalAssum (na,t) + | LocalDef (na,t',t) when is_topvar sigma (EConstr.of_constr t') -> LocalAssum (na,t) | _ -> decl) -let vars_of_ctx ctx = +let vars_of_ctx sigma ctx = let _, y = List.fold_right (fun decl (prev, vars) -> match decl with - | LocalDef (na,t',t) when is_topvar t' -> + | LocalDef (na,t',t) when is_topvar sigma (EConstr.of_constr t') -> prev, (GApp (Loc.ghost, (GRef (Loc.ghost, delayed_force coq_eq_refl_ref, None)), @@ -2213,12 +2269,12 @@ let constrs_of_pats typing_fun env evdref eqns tomatchs sign neqs arity = (* lift to get outside of past patterns to get terms in the combined environment. *) (fun (pats, n) (sign, c, (s, args), p) -> let len = List.length sign in - ((rels_of_patsign sign, lift n c, + ((rels_of_patsign !evdref sign, lift n c, (s, List.map (lift n) args), p) :: pats, len + n)) ([], 0) pats in let ineqs = build_ineqs evdref prevpatterns pats signlen in - let rhs_rels' = rels_of_patsign rhs_rels in + let rhs_rels' = rels_of_patsign !evdref rhs_rels in let _signenv = push_rel_context rhs_rels' env in let arity = let args, nargs = @@ -2234,21 +2290,21 @@ let constrs_of_pats typing_fun env evdref eqns tomatchs sign neqs arity = match ineqs with | None -> [], arity | Some ineqs -> - [LocalAssum (Anonymous, ineqs)], lift 1 arity + [local_assum (Anonymous, ineqs)], lift 1 arity in - let eqs_rels, arity = decompose_prod_n_assum neqs arity in + let eqs_rels, arity = decompose_prod_n_assum !evdref neqs arity in eqs_rels @ neqs_rels @ rhs_rels', arity in let rhs_env = push_rel_context rhs_rels' env in - let j = typing_fun (mk_tycon (EConstr.of_constr tycon)) rhs_env eqn.rhs.it in - let bbody = it_mkLambda_or_LetIn j.uj_val rhs_rels' - and btype = it_mkProd_or_LetIn j.uj_type rhs_rels' in - let _btype = evd_comb1 (Typing.type_of env) evdref (EConstr.of_constr bbody) in + let j = typing_fun (mk_tycon tycon) rhs_env eqn.rhs.it in + let bbody = it_mkLambda_or_LetIn (EConstr.of_constr j.uj_val) rhs_rels' + and btype = it_mkProd_or_LetIn (EConstr.of_constr j.uj_type) rhs_rels' in + let _btype = evd_comb1 (Typing.type_of env) evdref bbody in let branch_name = Id.of_string ("program_branch_" ^ (string_of_int !i)) in - let branch_decl = LocalDef (Name branch_name, lift !i bbody, lift !i btype) in + let branch_decl = local_def (Name branch_name, lift !i bbody, lift !i btype) in let branch = let bref = GVar (Loc.ghost, branch_name) in - match vars_of_ctx rhs_rels with + match vars_of_ctx !evdref rhs_rels with [] -> bref | l -> GApp (Loc.ghost, bref, l) in @@ -2287,14 +2343,14 @@ let abstract_tomatch env sigma tomatchs tycon = List.fold_left (fun (prev, ctx, names, tycon) (c, t) -> let lenctx = List.length ctx in - match kind_of_term c with + match EConstr.kind sigma c with Rel n -> (lift lenctx c, lift_tomatch_type lenctx t) :: prev, ctx, names, tycon | _ -> let tycon = Option.map - (fun t -> subst_term sigma (EConstr.of_constr (lift 1 c)) (EConstr.of_constr (lift 1 t))) tycon in + (fun t -> EConstr.of_constr (subst_term sigma (lift 1 c) (lift 1 t))) tycon in let name = next_ident_away (Id.of_string "filtered_var") names in (mkRel 1, lift_tomatch_type (succ lenctx) t) :: lift_ctx 1 prev, - LocalDef (Name name, lift lenctx c, lift lenctx $ type_of_tomatch t) :: ctx, + local_def (Name name, lift lenctx c, lift lenctx $ type_of_tomatch t) :: ctx, name :: names, tycon) ([], [], [], tycon) tomatchs in List.rev prev, ctx, tycon @@ -2315,21 +2371,25 @@ let build_dependent_signature env evdref avoid tomatchs arsign = *) match ty with | IsInd (ty, IndType (indf, args), _) when List.length args > 0 -> + let args = List.map EConstr.of_constr args in (* Build the arity signature following the names in matched terms as much as possible *) let argsign = List.tl arsign in (* arguments in inverse application order *) let app_decl = List.hd arsign in (* The matched argument *) let appn = RelDecl.get_name app_decl in let appt = RelDecl.get_type app_decl in + let appt = EConstr.of_constr appt in let argsign = List.rev argsign in (* arguments in application order *) let env', nargeqs, argeqs, refl_args, slift, argsign' = List.fold_left2 (fun (env, nargeqs, argeqs, refl_args, slift, argsign') arg decl -> let name = RelDecl.get_name decl in let t = RelDecl.get_type decl in - let argt = Retyping.get_type_of env !evdref (EConstr.of_constr arg) in + let t = EConstr.of_constr t in + let argt = Retyping.get_type_of env !evdref arg in + let argt = EConstr.of_constr argt in let eq, refl_arg = - if Reductionops.is_conv env !evdref (EConstr.of_constr argt) (EConstr.of_constr t) then + if Reductionops.is_conv env !evdref argt t then (mk_eq evdref (lift (nargeqs + slift) argt) (mkRel (nargeqs + slift)) (lift (nargeqs + nar) arg), @@ -2343,14 +2403,14 @@ let build_dependent_signature env evdref avoid tomatchs arsign = in let previd, id = let name = - match kind_of_term arg with + match EConstr.kind !evdref arg with Rel n -> RelDecl.get_name (lookup_rel n env) | _ -> name in make_prime avoid name in (env, succ nargeqs, - (LocalAssum (Name (eq_id avoid previd), eq)) :: argeqs, + (local_assum (Name (eq_id avoid previd), eq)) :: argeqs, refl_arg :: refl_args, pred slift, RelDecl.set_name (Name id) decl :: argsign')) @@ -2364,7 +2424,7 @@ let build_dependent_signature env evdref avoid tomatchs arsign = in let refl_eq = mk_JMeq_refl evdref ty tm in let previd, id = make_prime avoid appn in - ((LocalAssum (Name (eq_id avoid previd), eq) :: argeqs) :: eqs, + ((local_assum (Name (eq_id avoid previd), eq) :: argeqs) :: eqs, succ nargeqs, refl_eq :: refl_args, pred slift, @@ -2380,7 +2440,7 @@ let build_dependent_signature env evdref avoid tomatchs arsign = mk_eq evdref (lift nar tomatch_ty) (mkRel slift) (lift nar tm) in - ([LocalAssum (Name (eq_id avoid previd), eq)] :: eqs, succ neqs, + ([local_assum (Name (eq_id avoid previd), eq)] :: eqs, succ neqs, (mk_eq_refl evdref tomatch_ty tm) :: refl_args, pred slift, (arsign' :: []) :: arsigns)) ([], 0, [], nar, []) tomatchs arsign @@ -2409,6 +2469,7 @@ let compile_program_cases loc style (typing_function, evdref) tycon env (* constructors found in patterns *) let tomatchs = coerce_to_indtype typing_function evdref env matx tomatchl in let tycon = valcon_of_tycon tycon in + let tycon = Option.map EConstr.of_constr tycon in let tomatchs, tomatchs_lets, tycon' = abstract_tomatch env !evdref tomatchs tycon in let env = push_rel_context tomatchs_lets env in let len = List.length eqns in @@ -2454,9 +2515,9 @@ let compile_program_cases loc style (typing_function, evdref) tycon env (* We push the initial terms to match and push their alias to rhs' envs *) (* names of aliases will be recovered from patterns (hence Anonymous here) *) - let out_tmt na = function NotInd (None,t) -> LocalAssum (na,t) - | NotInd (Some b, t) -> LocalDef (na,b,t) - | IsInd (typ,_,_) -> LocalAssum (na,typ) in + let out_tmt na = function NotInd (None,t) -> local_assum (na,t) + | NotInd (Some b, t) -> local_def (na,b,t) + | IsInd (typ,_,_) -> local_assum (na,typ) in let typs = List.map2 (fun na (tm,tmt) -> (tm,out_tmt na tmt)) nal tomatchs in let typs = @@ -2470,7 +2531,7 @@ let compile_program_cases loc style (typing_function, evdref) tycon env let typs' = List.map3 (fun (tm,tmt) deps na -> - let deps = if not (isRel tm) then [] else deps in + let deps = if not (isRel !evdref tm) then [] else deps in ((tm,tmt),deps,na)) tomatchs dep_sign nal in @@ -2494,10 +2555,10 @@ let compile_program_cases loc style (typing_function, evdref) tycon env let j = compile pb in (* We check for unused patterns *) List.iter (check_unused_pattern env) matx; - let body = it_mkLambda_or_LetIn (applistc j.uj_val args) lets in + let body = it_mkLambda_or_LetIn (applist (EConstr.of_constr j.uj_val, args)) lets in let j = - { uj_val = it_mkLambda_or_LetIn body tomatchs_lets; - uj_type = nf_evar !evdref tycon; } + { uj_val = EConstr.Unsafe.to_constr (it_mkLambda_or_LetIn body tomatchs_lets); + uj_type = EConstr.to_constr !evdref tycon; } in j (**************************************************************************) @@ -2522,6 +2583,7 @@ let compile_cases loc style (typing_fun, evdref) tycon env (predopt, tomatchl, e with the type of arguments to match; if none is provided, we build alternative possible predicates *) let arsign = extract_arity_signature env tomatchs tomatchl in + let tycon = Option.map EConstr.of_constr tycon in let preds = prepare_predicate loc typing_fun env !evdref tomatchs arsign tycon predopt in let compile_for_one_predicate (sigma,nal,pred) = @@ -2529,9 +2591,9 @@ let compile_cases loc style (typing_fun, evdref) tycon env (predopt, tomatchl, e (* names of aliases will be recovered from patterns (hence Anonymous *) (* here) *) - let out_tmt na = function NotInd (None,t) -> LocalAssum (na,t) - | NotInd (Some b,t) -> LocalDef (na,b,t) - | IsInd (typ,_,_) -> LocalAssum (na,typ) in + let out_tmt na = function NotInd (None,t) -> local_assum (na,t) + | NotInd (Some b,t) -> local_def (na,b,t) + | IsInd (typ,_,_) -> local_assum (na,typ) in let typs = List.map2 (fun na (tm,tmt) -> (tm,out_tmt na tmt)) nal tomatchs in let typs = @@ -2545,7 +2607,7 @@ let compile_cases loc style (typing_fun, evdref) tycon env (predopt, tomatchl, e let typs' = List.map3 (fun (tm,tmt) deps na -> - let deps = if not (isRel tm) then [] else deps in + let deps = if not (isRel !evdref tm) then [] else deps in ((tm,tmt),deps,na)) tomatchs dep_sign nal in @@ -2572,7 +2634,7 @@ let compile_cases loc style (typing_fun, evdref) tycon env (predopt, tomatchl, e let j = compile pb in (* We coerce to the tycon (if an elim predicate was provided) *) - let j = inh_conv_coerce_to_tycon loc env myevdref j tycon in + let j = inh_conv_coerce_to_tycon loc env myevdref j (Option.map EConstr.Unsafe.to_constr tycon) in evdref := !myevdref; j in diff --git a/pretyping/cases.mli b/pretyping/cases.mli index 6bc61f6dd..9016ca5f3 100644 --- a/pretyping/cases.mli +++ b/pretyping/cases.mli @@ -8,6 +8,7 @@ open Names open Term +open EConstr open Evd open Environ open Inductiveops @@ -50,8 +51,8 @@ val constr_of_pat : Glob_term.cases_pattern -> Names.Id.t list -> Glob_term.cases_pattern * - (Context.Rel.Declaration.t list * Term.constr * - (Term.types * Term.constr list) * Glob_term.cases_pattern) * + (Context.Rel.Declaration.t list * constr * + (types * constr list) * Glob_term.cases_pattern) * Names.Id.t list type 'a rhs = @@ -117,7 +118,7 @@ val prepare_predicate : Loc.t -> Environ.env -> Evd.evar_map ref -> 'a -> Environ.unsafe_judgment) -> Environ.env -> Evd.evar_map -> - (Term.types * tomatch_type) list -> + (types * tomatch_type) list -> Context.Rel.t list -> - Constr.constr option -> - 'a option -> (Evd.evar_map * Names.name list * Term.constr) list + constr option -> + 'a option -> (Evd.evar_map * Names.name list * constr) list diff --git a/pretyping/indrec.ml b/pretyping/indrec.ml index 4025ca8b8..4fa5ad06d 100644 --- a/pretyping/indrec.ml +++ b/pretyping/indrec.ml @@ -382,9 +382,10 @@ let mis_make_indrec env sigma listdepkind mib u = arsign' in let obj = - Inductiveops.make_case_or_project env indf ci pred - (mkRel 1) branches + Inductiveops.make_case_or_project env !evdref indf ci (EConstr.of_constr pred) + (EConstr.mkRel 1) (Array.map EConstr.of_constr branches) in + let obj = EConstr.to_constr !evdref obj in it_mkLambda_or_LetIn_name env obj (Termops.lift_rel_context nrec deparsign) in diff --git a/pretyping/inductiveops.ml b/pretyping/inductiveops.ml index a9184777d..a93f2846b 100644 --- a/pretyping/inductiveops.ml +++ b/pretyping/inductiveops.ml @@ -343,24 +343,25 @@ let get_projections env (ind,params) = | Some (Some (id, projs, pbs)) -> Some projs | _ -> None -let make_case_or_project env indf ci pred c branches = +let make_case_or_project env sigma indf ci pred c branches = + let open EConstr in let projs = get_projections env indf in match projs with | None -> (mkCase (ci, pred, c, branches)) | Some ps -> assert(Array.length branches == 1); let () = - let _, _, t = destLambda pred in + let _, _, t = destLambda sigma pred in let (ind, _), _ = dest_ind_family indf in let mib, _ = Inductive.lookup_mind_specif env ind in - if (* dependent *) not (noccurn 1 t) && + if (* dependent *) not (Vars.noccurn sigma 1 t) && not (has_dependent_elim mib) then user_err ~hdr:"make_case_or_project" Pp.(str"Dependent case analysis not allowed" ++ str" on inductive type " ++ Names.MutInd.print (fst ind)) in let branch = branches.(0) in - let ctx, br = decompose_lam_n_assum (Array.length ps) branch in + let ctx, br = decompose_lam_n_assum sigma (Array.length ps) branch in let n, subst = List.fold_right (fun decl (i, subst) -> @@ -368,9 +369,9 @@ let make_case_or_project env indf ci pred c branches = | LocalAssum (na, t) -> let t = mkProj (Projection.make ps.(i) true, c) in (i + 1, t :: subst) - | LocalDef (na, b, t) -> (i, substl subst b :: subst)) + | LocalDef (na, b, t) -> (i, Vars.substl subst (EConstr.of_constr b) :: subst)) ctx (0, []) - in substl subst br + in Vars.substl subst br (* substitution in a signature *) diff --git a/pretyping/inductiveops.mli b/pretyping/inductiveops.mli index e375a2c6b..cf5523a50 100644 --- a/pretyping/inductiveops.mli +++ b/pretyping/inductiveops.mli @@ -185,8 +185,8 @@ val make_case_info : env -> inductive -> case_style -> case_info Fail with an error if the elimination is dependent while the inductive type does not allow dependent elimination. *) val make_case_or_project : - env -> inductive_family -> case_info -> - (* pred *) constr -> (* term *) constr -> (* branches *) constr array -> constr + env -> evar_map -> inductive_family -> case_info -> + (* pred *) EConstr.constr -> (* term *) EConstr.constr -> (* branches *) EConstr.constr array -> EConstr.constr (*i Compatibility val make_default_case_info : env -> case_style -> inductive -> case_info diff --git a/pretyping/program.ml b/pretyping/program.ml index 2606d91f3..8ec6083f7 100644 --- a/pretyping/program.ml +++ b/pretyping/program.ml @@ -58,7 +58,9 @@ let coq_JMeq_refl = init_reference ["Logic";"JMeq"] "JMeq_refl" let coq_not = init_constant ["Init";"Logic"] "not" let coq_and = init_constant ["Init";"Logic"] "and" -let mk_coq_not x = mkApp (delayed_force coq_not, [| x |]) +let delayed_force c = EConstr.of_constr (c ()) + +let mk_coq_not x = EConstr.mkApp (delayed_force coq_not, [| x |]) let unsafe_fold_right f = function hd :: tl -> List.fold_right f tl hd @@ -68,7 +70,7 @@ let mk_coq_and l = let and_typ = delayed_force coq_and in unsafe_fold_right (fun c conj -> - mkApp (and_typ, [| c ; conj |])) + EConstr.mkApp (and_typ, [| c ; conj |])) l (* true = transparent by default, false = opaque if possible *) diff --git a/pretyping/program.mli b/pretyping/program.mli index 64c4ca2c2..94a7bdcb6 100644 --- a/pretyping/program.mli +++ b/pretyping/program.mli @@ -6,7 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -open Term +open EConstr open Globnames (** A bunch of Coq constants used by Progam *) @@ -36,7 +36,7 @@ val mk_coq_and : constr list -> constr val mk_coq_not : constr -> constr (** Polymorphic application of delayed references *) -val papp : Evd.evar_map ref -> (unit -> global_reference) -> EConstr.constr array -> EConstr.constr +val papp : Evd.evar_map ref -> (unit -> global_reference) -> constr array -> constr val get_proofs_transparency : unit -> bool val is_program_cases : unit -> bool diff --git a/tactics/equality.ml b/tactics/equality.ml index 58c86ff42..9679ac402 100644 --- a/tactics/equality.ml +++ b/tactics/equality.ml @@ -864,7 +864,7 @@ let descend_then env sigma head dirn = List.map build_branch (List.interval 1 (Array.length mip.mind_consnames)) in let ci = make_case_info env ind RegularStyle in - Inductiveops.make_case_or_project env indf ci p head (Array.of_list brl))) + EConstr.Unsafe.to_constr (Inductiveops.make_case_or_project env sigma indf ci (EConstr.of_constr p) (EConstr.of_constr head) (Array.map_of_list EConstr.of_constr brl)))) (* Now we need to construct the discriminator, given a discriminable position. This boils down to: diff --git a/tactics/hipattern.ml b/tactics/hipattern.ml index d27e4afb7..87e252a38 100644 --- a/tactics/hipattern.ml +++ b/tactics/hipattern.ml @@ -110,7 +110,7 @@ let match_with_one_constructor sigma style onlybinary allow_rec t = Some (hdapp,args) else None else - let ctyp = prod_applist mip.mind_nf_lc.(0) args in + let ctyp = Term.prod_applist mip.mind_nf_lc.(0) args in let cargs = List.map RelDecl.get_type (prod_assum ctyp) in if not (is_lax_conjunction style) || has_nodep_prod sigma ctyp then (* Record or non strict conjunction *) @@ -176,7 +176,7 @@ let match_with_disjunction ?(strict=false) ?(onlybinary=false) sigma t = None else let cargs = - Array.map (fun ar -> pi2 (destProd (prod_applist ar args))) + Array.map (fun ar -> pi2 (destProd (Term.prod_applist ar args))) mip.mind_nf_lc in Some (hdapp,Array.to_list cargs) else diff --git a/toplevel/class.ml b/toplevel/class.ml index 46b212dee..6788cf1b7 100644 --- a/toplevel/class.ml +++ b/toplevel/class.ml @@ -198,7 +198,7 @@ let build_id_coercion idf_opt source poly = lams in let typ_f = - it_mkProd_wo_LetIn + List.fold_left (fun d c -> Term.mkProd_wo_LetIn c d) (mkProd (Anonymous, applistc vs (Context.Rel.to_extended_list 0 lams), lift 1 t)) lams in diff --git a/toplevel/himsg.ml b/toplevel/himsg.ml index df1f47e33..345eae9df 100644 --- a/toplevel/himsg.ml +++ b/toplevel/himsg.ml @@ -1202,6 +1202,7 @@ let explain_recursion_scheme_error = function (* Pattern-matching errors *) let explain_bad_pattern env sigma cstr ty = + let ty = EConstr.to_constr sigma ty in let env = make_all_name_different env in let pt = pr_lconstr_env env sigma ty in let pc = pr_constructor env cstr in @@ -1245,13 +1246,15 @@ let explain_non_exhaustive env pats = spc () ++ hov 0 (prlist_with_sep pr_comma pr_cases_pattern pats) let explain_cannot_infer_predicate env sigma typs = + let inj c = EConstr.to_constr sigma c in + let typs = Array.map_to_list (fun (c1, c2) -> (inj c1, inj c2)) typs in let env = make_all_name_different env in let pr_branch (cstr,typ) = let cstr,_ = decompose_app cstr in str "For " ++ pr_lconstr_env env sigma cstr ++ str ": " ++ pr_lconstr_env env sigma typ in str "Unable to unify the types found in the branches:" ++ - spc () ++ hov 0 (prlist_with_sep fnl pr_branch (Array.to_list typs)) + spc () ++ hov 0 (prlist_with_sep fnl pr_branch typs) let explain_pattern_matching_error env sigma = function | BadPattern (c,t) -> diff --git a/toplevel/indschemes.ml b/toplevel/indschemes.ml index fef99d8af..0723f16c3 100644 --- a/toplevel/indschemes.ml +++ b/toplevel/indschemes.ml @@ -488,7 +488,7 @@ let build_combined_scheme env schemes = let ctx, _ = list_split_rev_at prods (List.rev_map (fun (x, y) -> LocalAssum (x, y)) ctx) in - let typ = it_mkProd_wo_LetIn concl_typ ctx in + let typ = List.fold_left (fun d c -> Term.mkProd_wo_LetIn c d) concl_typ ctx in let body = it_mkLambda_or_LetIn concl_bod ctx in (body, typ) -- cgit v1.2.3 From 85ab3e298aa1d7333787c1fa44d25df189ac255c Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 8 Nov 2016 19:02:40 +0100 Subject: Pretyping API using EConstr. --- engine/evarutil.ml | 36 +++-- engine/evarutil.mli | 20 +-- engine/proofview.ml | 2 +- interp/constrintern.ml | 10 +- ltac/evar_tactics.ml | 2 +- ltac/extraargs.ml4 | 2 +- ltac/extratactics.ml4 | 2 +- ltac/rewrite.ml | 10 +- plugins/decl_mode/decl_interp.ml | 2 +- plugins/decl_mode/decl_proof_instr.ml | 2 +- pretyping/cases.ml | 6 +- pretyping/coercion.ml | 2 +- pretyping/evarconv.ml | 6 +- pretyping/evardefine.ml | 19 +-- pretyping/evardefine.mli | 15 +- pretyping/evarsolve.ml | 3 +- pretyping/nativenorm.ml | 2 - pretyping/pretyping.ml | 285 +++++++++++++++++++--------------- pretyping/pretyping.mli | 8 +- pretyping/reductionops.ml | 2 + pretyping/reductionops.mli | 8 +- pretyping/tacred.ml | 2 +- pretyping/unification.ml | 16 +- proofs/clenv.ml | 4 +- proofs/evar_refiner.ml | 2 +- proofs/logic.ml | 2 +- proofs/pfedit.ml | 2 +- proofs/pfedit.mli | 2 +- tactics/class_tactics.ml | 4 +- tactics/equality.ml | 2 +- tactics/hints.ml | 2 +- tactics/tactics.ml | 63 ++++---- toplevel/classes.ml | 6 +- toplevel/command.ml | 8 +- toplevel/record.ml | 6 +- toplevel/vernacentries.ml | 2 +- 36 files changed, 310 insertions(+), 257 deletions(-) diff --git a/engine/evarutil.ml b/engine/evarutil.ml index 62627a416..fc193b94a 100644 --- a/engine/evarutil.ml +++ b/engine/evarutil.ml @@ -278,17 +278,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 @@ -299,24 +302,29 @@ 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 * + csubst * (Id.t * EConstr.constr) list * Id.Set.t * Context.Named.t 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 EConstr in + let open Vars in + let map_decl f d = + NamedDecl.map_constr (fun c -> EConstr.Unsafe.to_constr (f (EConstr.of_constr c))) d + in let replace_var_named_declaration id0 id decl = 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 |> NamedDecl.set_id id' |> NamedDecl.map_constr (replace_vars vsubst) + decl |> NamedDecl.set_id id' |> map_decl (replace_vars vsubst) in let extract_if_neq id = function | Anonymous -> None @@ -344,7 +352,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 = decl |> NamedDecl.of_rel_decl (fun _ -> id0) |> NamedDecl.map_constr (subst2 subst vsubst) 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) | _ -> @@ -352,7 +360,7 @@ 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 = decl |> NamedDecl.of_rel_decl (fun _ -> id) |> NamedDecl.map_constr (subst2 subst vsubst) 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 = @@ -391,6 +399,7 @@ let new_pure_evar_full evd evi = Sigma.Unsafe.of_pair (evk, evd) let new_pure_evar sign evd ?(src=default_source) ?(filter = Filter.identity) ?candidates ?(store = Store.empty) ?naming ?(principal=false) typ = + let typ = EConstr.Unsafe.to_constr typ in let evd = Sigma.to_evar_map evd in let default_naming = Misctypes.IntroAnonymous in let naming = Option.default default_naming naming in @@ -420,7 +429,8 @@ let new_evar_instance sign evd typ ?src ?filter ?candidates ?store ?naming ?prin (* 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 map c = EConstr.Unsafe.to_constr (subst2 subst vsubst (EConstr.of_constr c)) in + let candidates = Option.map (fun l -> List.map map l) candidates in let instance = match filter with | None -> instance @@ -434,7 +444,7 @@ let new_evar_unsafe env evd ?src ?filter ?candidates ?store ?naming ?principal t 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 + let Sigma (e, evd', q) = new_evar env evd' ?src ?filter ?naming ?principal (EConstr.mkSort s) in Sigma ((e, s), evd', p +> q) let e_new_type_evar env evdref ?src ?filter ?naming ?principal rigid = @@ -753,5 +763,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 d6e2d4534..dcb9bf247 100644 --- a/engine/evarutil.mli +++ b/engine/evarutil.mli @@ -24,13 +24,13 @@ val new_evar : env -> 'r Sigma.t -> ?src:Loc.t * Evar_kinds.t -> ?filter:Filter.t -> ?candidates:constr list -> ?store:Store.t -> ?naming:Misctypes.intro_pattern_naming_expr -> - ?principal:bool -> types -> (constr, 'r) Sigma.sigma + ?principal:bool -> EConstr.types -> (constr, 'r) Sigma.sigma val new_pure_evar : named_context_val -> 'r Sigma.t -> ?src:Loc.t * Evar_kinds.t -> ?filter:Filter.t -> ?candidates:constr list -> ?store:Store.t -> ?naming:Misctypes.intro_pattern_naming_expr -> - ?principal:bool -> types -> (evar, 'r) Sigma.sigma + ?principal:bool -> EConstr.types -> (evar, 'r) Sigma.sigma val new_pure_evar_full : 'r Sigma.t -> evar_info -> (evar, 'r) Sigma.sigma @@ -39,7 +39,7 @@ val e_new_evar : env -> evar_map ref -> ?src:Loc.t * Evar_kinds.t -> ?filter:Filter.t -> ?candidates:constr list -> ?store:Store.t -> ?naming:Misctypes.intro_pattern_naming_expr -> - ?principal:bool -> types -> constr + ?principal:bool -> EConstr.types -> constr (** Create a new Type existential variable, as we keep track of them during type-checking and unification. *) @@ -70,7 +70,7 @@ 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 -> + named_context_val -> 'r Sigma.t -> EConstr.types -> ?src:Loc.t * Evar_kinds.t -> ?filter:Filter.t -> ?candidates:constr list -> ?store:Store.t -> ?naming:Misctypes.intro_pattern_naming_expr -> ?principal:bool -> @@ -208,17 +208,17 @@ 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 -> EConstr.t -> EConstr.t type ext_named_context = - csubst * (Id.t * Constr.constr) list * + csubst * (Id.t * EConstr.constr) list * Id.Set.t * Context.Named.t val push_rel_decl_to_named_context : Context.Rel.Declaration.t -> ext_named_context -> ext_named_context -val push_rel_context_to_named_context : Environ.env -> types -> - named_context_val * types * constr list * csubst * (identifier*constr) list +val push_rel_context_to_named_context : Environ.env -> EConstr.types -> + named_context_val * EConstr.types * constr list * csubst * (identifier*EConstr.constr) list val generalize_evar_over_rels : evar_map -> existential -> types * constr list @@ -235,5 +235,5 @@ val meta_counter_summary_name : string (** Deprecater *) -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/proofview.ml b/engine/proofview.ml index 21227ed19..b0f6d463b 100644 --- a/engine/proofview.ml +++ b/engine/proofview.ml @@ -71,7 +71,7 @@ let dependent_init = | 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 (econstr, sigma, _) = Evarutil.new_evar env sigma ~src ~store (EConstr.of_constr typ) in let sigma = Sigma.to_evar_map sigma in let ret, { solution = sol; comb = comb } = aux (t sigma econstr) in let (gl, _) = Term.destEvar econstr in diff --git a/interp/constrintern.ml b/interp/constrintern.ml index 235e6e24f..8d4d87f2a 100644 --- a/interp/constrintern.ml +++ b/interp/constrintern.ml @@ -1937,7 +1937,7 @@ let extract_ids env = let scope_of_type_kind = function | IsType -> Notation.current_type_scope_name () - | OfType typ -> compute_type_scope typ + | OfType typ -> compute_type_scope (EConstr.Unsafe.to_constr typ) | WithoutTypeConstraint -> None let empty_ltac_sign = { @@ -1982,7 +1982,7 @@ let interp_type env sigma ?(impls=empty_internalization_env) c = interp_gen IsType env sigma ~impls c let interp_casted_constr env sigma ?(impls=empty_internalization_env) c typ = - interp_gen (OfType typ) env sigma ~impls c + interp_gen (OfType (EConstr.of_constr typ)) env sigma ~impls c (* Not all evars expected to be resolved *) @@ -2001,7 +2001,7 @@ let interp_constr_evars_impls env evdref ?(impls=empty_internalization_env) c = interp_constr_evars_gen_impls env evdref ~impls WithoutTypeConstraint c let interp_casted_constr_evars_impls env evdref ?(impls=empty_internalization_env) c typ = - interp_constr_evars_gen_impls env evdref ~impls (OfType typ) c + interp_constr_evars_gen_impls env evdref ~impls (OfType (EConstr.of_constr typ)) c let interp_type_evars_impls env evdref ?(impls=empty_internalization_env) c = interp_constr_evars_gen_impls env evdref ~impls IsType c @@ -2016,7 +2016,7 @@ let interp_constr_evars env evdref ?(impls=empty_internalization_env) c = interp_constr_evars_gen env evdref WithoutTypeConstraint ~impls c let interp_casted_constr_evars env evdref ?(impls=empty_internalization_env) c typ = - interp_constr_evars_gen env evdref ~impls (OfType typ) c + interp_constr_evars_gen env evdref ~impls (OfType (EConstr.of_constr typ)) c let interp_type_evars env evdref ?(impls=empty_internalization_env) c = interp_constr_evars_gen env evdref IsType ~impls c @@ -2102,7 +2102,7 @@ let interp_rawcontext_evars env evdref k bl = in (push_rel d env, d::params, succ n, impls) | Some b -> - let c = understand_tcc_evars env evdref ~expected_type:(OfType t) b in + let c = understand_tcc_evars env evdref ~expected_type:(OfType (EConstr.of_constr t)) b in let d = LocalDef (na, c, t) in (push_rel d env, d::params, n, impls)) (env,[],k+1,[]) (List.rev bl) diff --git a/ltac/evar_tactics.ml b/ltac/evar_tactics.ml index 5d3b2b886..99023fdbb 100644 --- a/ltac/evar_tactics.ml +++ b/ltac/evar_tactics.ml @@ -85,7 +85,7 @@ let let_evar name typ = Namegen.next_ident_away_in_goal id (Termops.ids_of_named_context (Environ.named_context env)) | Names.Name id -> id in - let Sigma (evar, sigma, p) = Evarutil.new_evar env sigma ~src ~naming:(Misctypes.IntroFresh id) typ in + let Sigma (evar, sigma, p) = Evarutil.new_evar env sigma ~src ~naming:(Misctypes.IntroFresh id) (EConstr.of_constr typ) in let tac = (Tactics.letin_tac None (Names.Name id) evar None Locusops.nowhere) in diff --git a/ltac/extraargs.ml4 b/ltac/extraargs.ml4 index 53b726432..f8db0b4fc 100644 --- a/ltac/extraargs.ml4 +++ b/ltac/extraargs.ml4 @@ -177,7 +177,7 @@ ARGUMENT EXTEND lglob END let interp_casted_constr ist gl c = - interp_constr_gen (Pretyping.OfType (pf_concl gl)) ist (pf_env gl) (project gl) c + interp_constr_gen (Pretyping.OfType (EConstr.of_constr (pf_concl gl))) ist (pf_env gl) (project gl) c ARGUMENT EXTEND casted_constr TYPED AS constr diff --git a/ltac/extratactics.ml4 b/ltac/extratactics.ml4 index e1b468197..981ff549d 100644 --- a/ltac/extratactics.ml4 +++ b/ltac/extratactics.ml4 @@ -351,7 +351,7 @@ let refine_tac ist simple c = let concl = Proofview.Goal.concl gl in let env = Proofview.Goal.env gl in let flags = constr_flags in - let expected_type = Pretyping.OfType concl in + let expected_type = Pretyping.OfType (EConstr.of_constr concl) in let c = Pretyping.type_uconstr ~flags ~expected_type ist c in let update = { run = fun sigma -> c.delayed env sigma } in let refine = Refine.refine ~unsafe:true update in diff --git a/ltac/rewrite.ml b/ltac/rewrite.ml index 58153c453..ccd45756a 100644 --- a/ltac/rewrite.ml +++ b/ltac/rewrite.ml @@ -95,7 +95,7 @@ let cstrevars evars = snd evars let new_cstr_evar (evd,cstrs) env t = let s = Typeclasses.set_resolvable Evd.Store.empty false in let evd = Sigma.Unsafe.of_evar_map evd in - let Sigma (t, evd', _) = Evarutil.new_evar ~store:s env evd t in + let Sigma (t, evd', _) = Evarutil.new_evar ~store:s env evd (EConstr.of_constr t) in let evd' = Sigma.to_evar_map evd' in let ev, _ = destEvar t in (evd', Evar.Set.add ev cstrs), t @@ -1549,8 +1549,8 @@ let assert_replacing id newt tac = in let env' = Environ.reset_with_named_context (val_of_named_context nc) env in Refine.refine ~unsafe:false { run = begin fun sigma -> - let Sigma (ev, sigma, p) = Evarutil.new_evar env' sigma concl in - let Sigma (ev', sigma, q) = Evarutil.new_evar env sigma newt in + let Sigma (ev, sigma, p) = Evarutil.new_evar env' sigma (EConstr.of_constr concl) in + let Sigma (ev', sigma, q) = Evarutil.new_evar env sigma (EConstr.of_constr newt) in let map d = let n = NamedDecl.get_id d in if Id.equal n id then ev' else mkVar n @@ -1596,7 +1596,7 @@ let cl_rewrite_clause_newtac ?abs ?origsigma ~progress strat clause = Proofview.Goal.enter { enter = begin fun gl -> let env = Proofview.Goal.env gl in let make = { run = begin fun sigma -> - let Sigma (ev, sigma, q) = Evarutil.new_evar env sigma newt in + let Sigma (ev, sigma, q) = Evarutil.new_evar env sigma (EConstr.of_constr newt) in Sigma (mkApp (p, [| ev |]), sigma, q) end } in Refine.refine ~unsafe:false make <*> Proofview.Unsafe.tclNEWGOALS gls @@ -1926,7 +1926,7 @@ let build_morphism_signature env sigma m = let evd = solve_constraints env !evd in let evd = Evd.nf_constraints evd in let m = Evarutil.nf_evars_universes evd morph in - Pretyping.check_evars env Evd.empty evd m; + Pretyping.check_evars env Evd.empty evd (EConstr.of_constr m); Evd.evar_universe_context evd, m let default_morphism sign m = diff --git a/plugins/decl_mode/decl_interp.ml b/plugins/decl_mode/decl_interp.ml index 65d273faf..ddf013735 100644 --- a/plugins/decl_mode/decl_interp.ml +++ b/plugins/decl_mode/decl_interp.ml @@ -173,7 +173,7 @@ let get_eq_typ info env = typ let interp_constr_in_type typ env sigma c = - fst (understand env sigma (fst c) ~expected_type:(OfType typ))(*FIXME*) + fst (understand env sigma (fst c) ~expected_type:(OfType (EConstr.of_constr typ)))(*FIXME*) let interp_statement interp_it env sigma st = {st_label=st.st_label; diff --git a/plugins/decl_mode/decl_proof_instr.ml b/plugins/decl_mode/decl_proof_instr.ml index e587fd52c..5e16d2da0 100644 --- a/plugins/decl_mode/decl_proof_instr.ml +++ b/plugins/decl_mode/decl_proof_instr.ml @@ -1330,7 +1330,7 @@ let understand_my_constr env sigma c concl = | GEvar _ -> GHole (Loc.ghost,Evar_kinds.QuestionMark Evar_kinds.Expand,Misctypes.IntroAnonymous,None) | rc -> map_glob_constr frob rc in - Pretyping.understand_tcc env sigma ~expected_type:(Pretyping.OfType concl) (frob rawc) + Pretyping.understand_tcc env sigma ~expected_type:(Pretyping.OfType (EConstr.of_constr concl)) (frob rawc) let my_refine c gls = let oc = { run = begin fun sigma -> diff --git a/pretyping/cases.ml b/pretyping/cases.ml index 1a181202c..92bd1e389 100644 --- a/pretyping/cases.ml +++ b/pretyping/cases.ml @@ -301,7 +301,7 @@ let inductive_template evdref env tmloc ind = | LocalAssum (na,ty) -> let ty = EConstr.of_constr ty in let ty' = substl subst ty in - let e = EConstr.of_constr (e_new_evar env evdref ~src:(hole_source n) (EConstr.Unsafe.to_constr ty')) in + let e = EConstr.of_constr (e_new_evar env evdref ~src:(hole_source n) ty') in (e::subst,e::evarl,n+1) | LocalDef (na,b,ty) -> let b = EConstr.of_constr b in @@ -1665,6 +1665,7 @@ let abstract_tycon loc env evdref subst tycon extenv t = (fun i _ -> try list_assoc_in_triple i subst0 with Not_found -> mkRel i) 1 (rel_context env) in + let ty = EConstr.of_constr ty in let ev' = e_new_evar env evdref ~src ty in let ev' = EConstr.of_constr ev' in begin match solve_simple_eqn (evar_conv_x full_transparent_state) env !evdref (None,ev,substl inst ev') with @@ -1700,7 +1701,6 @@ let abstract_tycon loc env evdref subst tycon extenv t = let filter = Filter.make (rel_filter @ named_filter) in let candidates = u :: List.map mkRel vl in let candidates = List.map EConstr.Unsafe.to_constr candidates in - let ty = EConstr.Unsafe.to_constr ty in let ev = e_new_evar extenv evdref ~src ~filter ~candidates ty in let ev = EConstr.of_constr ev in lift k ev @@ -2469,7 +2469,6 @@ let compile_program_cases loc style (typing_function, evdref) tycon env (* constructors found in patterns *) let tomatchs = coerce_to_indtype typing_function evdref env matx tomatchl in let tycon = valcon_of_tycon tycon in - let tycon = Option.map EConstr.of_constr tycon in let tomatchs, tomatchs_lets, tycon' = abstract_tomatch env !evdref tomatchs tycon in let env = push_rel_context tomatchs_lets env in let len = List.length eqns in @@ -2583,7 +2582,6 @@ let compile_cases loc style (typing_fun, evdref) tycon env (predopt, tomatchl, e with the type of arguments to match; if none is provided, we build alternative possible predicates *) let arsign = extract_arity_signature env tomatchs tomatchl in - let tycon = Option.map EConstr.of_constr tycon in let preds = prepare_predicate loc typing_fun env !evdref tomatchs arsign tycon predopt in let compile_for_one_predicate (sigma,nal,pred) = diff --git a/pretyping/coercion.ml b/pretyping/coercion.ml index cc121a96d..b9f14aa43 100644 --- a/pretyping/coercion.ml +++ b/pretyping/coercion.ml @@ -93,7 +93,7 @@ open Program let make_existential loc ?(opaque = not (get_proofs_transparency ())) env evdref c = let src = (loc, Evar_kinds.QuestionMark (Evar_kinds.Define opaque)) in - EConstr.of_constr (Evarutil.e_new_evar env evdref ~src (EConstr.Unsafe.to_constr c)) + EConstr.of_constr (Evarutil.e_new_evar env evdref ~src c) let app_opt env evdref f t = EConstr.of_constr (whd_betaiota !evdref (app_opt f t)) diff --git a/pretyping/evarconv.ml b/pretyping/evarconv.ml index cdcb993b5..683b33b89 100644 --- a/pretyping/evarconv.ml +++ b/pretyping/evarconv.ml @@ -338,7 +338,7 @@ let rec evar_conv_x ts env evd pbty term1 term2 = let e = try let evd, b = infer_conv ~catch_incon:false ~pb:pbty ~ts:(fst ts) - env evd (EConstr.Unsafe.to_constr term1) (EConstr.Unsafe.to_constr term2) + env evd term1 term2 in if b then Success evd else UnifFailure (evd, ConversionFailed (env,term1,term2)) @@ -891,7 +891,7 @@ and conv_record trs env evd (ctx,(h,h2),c,bs,(params,params1),(us,us2),(sk1,sk2) else let dloc = (Loc.ghost,Evar_kinds.InternalHole) in let i = Sigma.Unsafe.of_evar_map i in - let Sigma (ev, i', _) = Evarutil.new_evar env i ~src:dloc (EConstr.Unsafe.to_constr (Vars.substl ks b)) in + let Sigma (ev, i', _) = Evarutil.new_evar env i ~src:dloc (Vars.substl ks b) in let i' = Sigma.to_evar_map i' in (i', EConstr.of_constr ev :: ks, m - 1,test)) (evd,[],List.length bs,fun i -> Success i) bs @@ -1075,7 +1075,7 @@ let second_order_matching ts env_rhs evd (evk,args) argoccs rhs = let evty = set_holes evdref cty subst in let instance = List.map EConstr.Unsafe.to_constr (Filter.filter_list filter instance) in let evd = Sigma.Unsafe.of_evar_map !evdref in - let Sigma (ev, evd, _) = new_evar_instance sign evd (EConstr.Unsafe.to_constr evty) ~filter instance in + let Sigma (ev, evd, _) = new_evar_instance sign evd evty ~filter instance in let evd = Sigma.to_evar_map evd in evdref := evd; evsref := (fst (destEvar !evdref (EConstr.of_constr ev)),evty)::!evsref; diff --git a/pretyping/evardefine.ml b/pretyping/evardefine.ml index f372dbf06..ff40a6938 100644 --- a/pretyping/evardefine.ml +++ b/pretyping/evardefine.ml @@ -39,9 +39,9 @@ let env_nf_betaiotaevar sigma env = (* Operations on value/type constraints *) (****************************************) -type type_constraint = types option +type type_constraint = EConstr.types option -type val_constraint = constr option +type val_constraint = EConstr.constr option (* Old comment... * Basically, we have the following kind of constraints (in increasing @@ -61,13 +61,13 @@ type val_constraint = constr option let empty_tycon = None (* Builds a type constraint *) -let mk_tycon ty = Some (EConstr.Unsafe.to_constr ty) +let mk_tycon ty = Some ty (* Constrains the value of a type *) let empty_valcon = None (* Builds a value constraint *) -let mk_valcon c = Some (EConstr.Unsafe.to_constr c) +let mk_valcon c = Some c let idx = Namegen.default_dependent_ident @@ -80,7 +80,8 @@ let define_pure_evar_as_product evd evk = let evenv = evar_env evi in let id = next_ident_away idx (ids_of_named_context (evar_context evi)) in let concl = Reductionops.whd_all evenv evd (EConstr.of_constr evi.evar_concl) in - let s = destSort evd (EConstr.of_constr concl) in + let concl = EConstr.of_constr concl in + let s = destSort evd concl in let evd1,(dom,u1) = let evd = Sigma.Unsafe.of_evar_map evd in let Sigma (e, evd1, _) = new_type_evar evenv evd univ_flexible_alg ~filter:(evar_filter evi) in @@ -146,7 +147,7 @@ let define_pure_evar_as_lambda env evd evk = let newenv = push_named (LocalAssum (id, dom)) evenv in let filter = Filter.extend 1 (evar_filter evi) in let src = evar_source evk evd1 in - let evd2,body = new_evar_unsafe newenv evd1 ~src (EConstr.Unsafe.to_constr (Vars.subst1 (mkVar id) rng)) ~filter in + let evd2,body = new_evar_unsafe newenv evd1 ~src (Vars.subst1 (mkVar id) rng) ~filter in let lam = mkLambda (Name id, EConstr.of_constr dom, Vars.subst_var id (EConstr.of_constr body)) in Evd.define evk (EConstr.Unsafe.to_constr lam) evd2, lam @@ -203,12 +204,12 @@ let split_tycon loc env evd tycon = match tycon with | None -> evd,(Anonymous,None,None) | Some c -> - let evd', (n, dom, rng) = real_split evd (EConstr.of_constr c) in + let evd', (n, dom, rng) = real_split evd c in evd', (n, mk_tycon dom, mk_tycon rng) let valcon_of_tycon x = x -let lift_tycon n = Option.map (lift n) +let lift_tycon n = Option.map (EConstr.Vars.lift n) let pr_tycon env = function None -> str "None" - | Some t -> Termops.print_constr_env env t + | Some t -> Termops.print_constr_env env (EConstr.Unsafe.to_constr t) diff --git a/pretyping/evardefine.mli b/pretyping/evardefine.mli index f7bf4636b..9c03a6e3f 100644 --- a/pretyping/evardefine.mli +++ b/pretyping/evardefine.mli @@ -8,6 +8,7 @@ open Names open Term +open EConstr open Evd open Environ @@ -18,16 +19,16 @@ type type_constraint = types option type val_constraint = constr option val empty_tycon : type_constraint -val mk_tycon : EConstr.constr -> type_constraint +val mk_tycon : constr -> type_constraint val empty_valcon : val_constraint -val mk_valcon : EConstr.constr -> val_constraint +val mk_valcon : constr -> val_constraint (** Instantiate an evar by as many lambda's as needed so that its arguments are moved to the evar substitution (i.e. turn [?x[vars1:=args1] args] into [?y[vars1:=args1,vars:=args]] with [vars1 |- ?x:=\vars.?y[vars1:=vars1,vars:=vars]] *) -val evar_absorb_arguments : env -> evar_map -> EConstr.existential -> EConstr.constr list -> - evar_map * EConstr.existential +val evar_absorb_arguments : env -> evar_map -> existential -> constr list -> + evar_map * existential val split_tycon : Loc.t -> env -> evar_map -> type_constraint -> @@ -36,9 +37,9 @@ val split_tycon : val valcon_of_tycon : type_constraint -> val_constraint val lift_tycon : int -> type_constraint -> type_constraint -val define_evar_as_product : evar_map -> EConstr.existential -> evar_map * EConstr.types -val define_evar_as_lambda : env -> evar_map -> EConstr.existential -> evar_map * EConstr.types -val define_evar_as_sort : env -> evar_map -> EConstr.existential -> evar_map * sorts +val define_evar_as_product : evar_map -> existential -> evar_map * types +val define_evar_as_lambda : env -> evar_map -> existential -> evar_map * types +val define_evar_as_sort : env -> evar_map -> existential -> evar_map * sorts (** {6 debug pretty-printer:} *) diff --git a/pretyping/evarsolve.ml b/pretyping/evarsolve.ml index 8a22aed2f..3bcea4cee 100644 --- a/pretyping/evarsolve.ml +++ b/pretyping/evarsolve.ml @@ -606,7 +606,7 @@ let make_projectable_subst aliases sigma evi args = let define_evar_from_virtual_equation define_fun env evd src t_in_env ty_t_in_sign sign filter inst_in_env = let open EConstr in let evd = Sigma.Unsafe.of_evar_map evd in - let Sigma (evar_in_env, evd, _) = new_evar_instance sign evd ty_t_in_sign ~filter ~src (List.map EConstr.Unsafe.to_constr inst_in_env) in + let Sigma (evar_in_env, evd, _) = new_evar_instance sign evd (EConstr.of_constr ty_t_in_sign) ~filter ~src (List.map EConstr.Unsafe.to_constr inst_in_env) in let evd = Sigma.to_evar_map evd in let t_in_env = EConstr.of_constr (whd_evar evd (EConstr.Unsafe.to_constr t_in_env)) in let evar_in_env = EConstr.of_constr evar_in_env in @@ -682,6 +682,7 @@ let materialize_evar define_fun env evd k (evk1,args1) ty_in_env = define_evar_from_virtual_equation define_fun env evd src ty_in_env ty_t_in_sign sign2 filter2 inst2_in_env in let evd = Sigma.Unsafe.of_evar_map evd in + let ev2ty_in_sign = EConstr.of_constr ev2ty_in_sign in let Sigma (ev2_in_sign, evd, _) = new_evar_instance sign2 evd ev2ty_in_sign ~filter:filter2 ~src (List.map EConstr.Unsafe.to_constr inst2_in_sign) in let evd = Sigma.to_evar_map evd in diff --git a/pretyping/nativenorm.ml b/pretyping/nativenorm.ml index c8bcae0c8..ff3424c44 100644 --- a/pretyping/nativenorm.ml +++ b/pretyping/nativenorm.ml @@ -409,7 +409,5 @@ let native_conv_generic pb sigma t = Nativeconv.native_conv_gen pb (evars_of_evar_map sigma) t let native_infer_conv ?(pb=Reduction.CUMUL) env sigma t1 t2 = - let t1 = EConstr.Unsafe.to_constr t1 in - let t2 = EConstr.Unsafe.to_constr t2 in Reductionops.infer_conv_gen (fun pb ~l2r sigma ts -> native_conv_generic pb sigma) ~catch_incon:true ~pb env sigma t1 t2 diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml index 18731f1e9..cac31a1c5 100644 --- a/pretyping/pretyping.ml +++ b/pretyping/pretyping.ml @@ -27,8 +27,9 @@ open Util open Names open Evd open Term -open Vars open Termops +open EConstr +open Vars open Reductionops open Environ open Type_errors @@ -59,7 +60,7 @@ type ltac_var_map = { ltac_genargs : unbound_ltac_var_map; } type glob_constr_ltac_closure = ltac_var_map * glob_constr -type pure_open_constr = evar_map * constr +type pure_open_constr = evar_map * Constr.constr (************************************************************************) (* This concerns Cases *) @@ -68,6 +69,16 @@ open Inductiveops (************************************************************************) +let local_assum (na, t) = + let open Context.Rel.Declaration in + let inj = EConstr.Unsafe.to_constr in + LocalAssum (na, inj t) + +let local_def (na, b, t) = + let open Context.Rel.Declaration in + let inj = EConstr.Unsafe.to_constr in + LocalDef (na, inj b, inj t) + module ExtraEnv = struct @@ -104,7 +115,7 @@ let lookup_named id env = lookup_named id env.env let e_new_evar env evdref ?src ?naming typ = let subst2 subst vsubst c = csubst_subst subst (replace_vars vsubst c) in let open Context.Named.Declaration in - let inst_vars = List.map (get_id %> mkVar) (named_context env.env) in + let inst_vars = List.map (get_id %> Constr.mkVar) (named_context env.env) in let inst_rels = List.rev (rel_list 0 (nb_rel env.env)) in let (subst, vsubst, _, nc) = Lazy.force env.extra in let typ' = subst2 subst vsubst typ in @@ -116,7 +127,7 @@ let e_new_evar env evdref ?src ?naming typ = e let push_rec_types (lna,typarray,_) env = - let ctxt = Array.map2_i (fun i na t -> Context.Rel.Declaration.LocalAssum (na, lift i t)) lna typarray in + let ctxt = Array.map2_i (fun i na t -> local_assum (na, lift i t)) lna typarray in Array.fold_left (fun e assum -> push_rel assum e) env ctxt end @@ -127,6 +138,13 @@ open ExtraEnv exception Found of int array +let nf_fix sigma (nas, cs, ts) = + let inj c = EConstr.to_constr sigma c in + (nas, Array.map inj cs, Array.map inj ts) + +let nf_evar sigma c = + EConstr.of_constr (nf_evar sigma (EConstr.Unsafe.to_constr c)) + let search_guard loc env possible_indexes fixdefs = (* Standard situation with only one possibility for each fix. *) (* We treat it separately in order to get proper error msg. *) @@ -282,7 +300,7 @@ let apply_inference_hook hook evdref pending = then try let sigma, c = hook sigma evk in - Evd.define evk c sigma + Evd.define evk (EConstr.Unsafe.to_constr c) sigma with Exit -> sigma else @@ -313,17 +331,15 @@ let check_extra_evars_are_solved env current_sigma pending = let check_evars env initial_sigma sigma c = let rec proc_rec c = - match kind_of_term c with - | Evar (evk,_ as ev) -> - (match existential_opt_value sigma ev with - | Some c -> proc_rec c - | None -> - if not (Evd.mem initial_sigma evk) then - let (loc,k) = evar_source evk sigma in - match k with - | Evar_kinds.ImplicitArg (gr, (i, id), false) -> () - | _ -> Pretype_errors.error_unsolvable_implicit ~loc env sigma evk None) - | _ -> Constr.iter proc_rec c + match EConstr.kind sigma c with + | Evar (evk, _) -> + if not (Evd.mem initial_sigma evk) then + let (loc,k) = evar_source evk sigma in + begin match k with + | Evar_kinds.ImplicitArg (gr, (i, id), false) -> () + | _ -> Pretype_errors.error_unsolvable_implicit ~loc env sigma evk None + end + | _ -> EConstr.iter sigma proc_rec c in proc_rec c let check_evars_are_solved env current_sigma frozen pending = @@ -359,7 +375,7 @@ let allow_anonymous_refs = ref false let inh_conv_coerce_to_tycon resolve_tc loc env evdref j = function | None -> j | Some t -> - evd_comb2 (Coercion.inh_conv_coerce_to resolve_tc loc env.ExtraEnv.env) evdref j (EConstr.of_constr t) + evd_comb2 (Coercion.inh_conv_coerce_to resolve_tc loc env.ExtraEnv.env) evdref j t let check_instance loc subst = function | [] -> () @@ -409,26 +425,29 @@ let invert_ltac_bound_name lvar env id0 id = str " which is not bound in current context.") let protected_get_type_of env sigma c = - try Retyping.get_type_of ~lax:true env.ExtraEnv.env sigma (EConstr.of_constr c) + try Retyping.get_type_of ~lax:true env.ExtraEnv.env sigma c with Retyping.RetypeError _ -> user_err - (str "Cannot reinterpret " ++ quote (print_constr c) ++ + (str "Cannot reinterpret " ++ quote (print_constr (EConstr.Unsafe.to_constr c)) ++ str " in the current environment.") +let j_val j = EConstr.of_constr (j_val j) + let pretype_id pretype k0 loc env evdref lvar id = let sigma = !evdref in (* Look for the binder of [id] *) try let (n,_,typ) = lookup_rel_id id (rel_context env) in - { uj_val = mkRel n; uj_type = lift n typ } + let typ = EConstr.of_constr typ in + { uj_val = EConstr.Unsafe.to_constr (mkRel n); uj_type = EConstr.Unsafe.to_constr (lift n typ) } with Not_found -> let env = ltac_interp_name_env k0 lvar env in (* Check if [id] is an ltac variable *) try let (ids,c) = Id.Map.find id lvar.ltac_constrs in let subst = List.map (invert_ltac_bound_name lvar env id) ids in - let c = substl subst c in - { uj_val = c; uj_type = protected_get_type_of env sigma c } + let c = substl subst (EConstr.of_constr c) in + { uj_val = EConstr.Unsafe.to_constr c; uj_type = protected_get_type_of env sigma c } with Not_found -> try let {closure;term} = Id.Map.find id lvar.ltac_uconstrs in let lvar = { @@ -453,14 +472,11 @@ let pretype_id pretype k0 loc env evdref lvar id = end; (* Check if [id] is a section or goal variable *) try - { uj_val = mkVar id; uj_type = NamedDecl.get_type (lookup_named id env) } + { uj_val = Constr.mkVar id; uj_type = NamedDecl.get_type (lookup_named id env) } with Not_found -> (* [id] not found, standard error message *) error_var_not_found ~loc id -let evar_kind_of_term sigma c = - kind_of_term (whd_evar sigma c) - (*************************************************************************) (* Main pretyping function *) @@ -492,13 +508,17 @@ let pretype_global loc rigid env evd gr us = str " universe instances must be greater or equal to Set."); evd, Some (Univ.Instance.of_array (Array.of_list (List.rev l'))) in - Evd.fresh_global ~loc ~rigid ?names:instance env.ExtraEnv.env evd gr + let (sigma, c) = Evd.fresh_global ~loc ~rigid ?names:instance env.ExtraEnv.env evd gr in + (sigma, EConstr.of_constr c) + +let make_judge c t = + make_judge (EConstr.Unsafe.to_constr c) (EConstr.Unsafe.to_constr t) let pretype_ref loc evdref env ref us = match ref with | VarRef id -> (* Section variable *) - (try make_judge (mkVar id) (NamedDecl.get_type (lookup_named id env)) + (try make_judge (mkVar id) (EConstr.of_constr (NamedDecl.get_type (lookup_named id env))) with Not_found -> (* This may happen if env is a goal env and section variables have been cleared - section variables should be different from goal @@ -507,13 +527,14 @@ let pretype_ref loc evdref env ref us = | ref -> let evd, c = pretype_global loc univ_flexible env !evdref ref us in let () = evdref := evd in - let ty = Typing.unsafe_type_of env.ExtraEnv.env evd (EConstr.of_constr c) in + let ty = Typing.unsafe_type_of env.ExtraEnv.env evd c in + let ty = EConstr.of_constr ty in make_judge c ty let judge_of_Type loc evd s = let evd, s = interp_universe ~loc evd s in let judge = - { uj_val = mkSort (Type s); uj_type = mkSort (Type (Univ.super s)) } + { uj_val = Constr.mkSort (Type s); uj_type = Constr.mkSort (Type (Univ.super s)) } in evd, judge @@ -563,32 +584,32 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre let hyps = evar_filtered_context (Evd.find !evdref evk) in let args = pretype_instance k0 resolve_tc env evdref lvar loc hyps evk inst in let c = mkEvar (evk, args) in - let j = (Retyping.get_judgment_of env.ExtraEnv.env !evdref (EConstr.of_constr c)) in + let j = (Retyping.get_judgment_of env.ExtraEnv.env !evdref c) in inh_conv_coerce_to_tycon loc env evdref j tycon | GPatVar (loc,(someta,n)) -> let env = ltac_interp_name_env k0 lvar env in let ty = match tycon with - | Some ty -> ty + | Some ty -> EConstr.Unsafe.to_constr ty | None -> new_type_evar env evdref loc in let k = Evar_kinds.MatchingVar (someta,n) in - { uj_val = e_new_evar env evdref ~src:(loc,k) ty; uj_type = ty } + { uj_val = e_new_evar env evdref ~src:(loc,k) (EConstr.of_constr ty); uj_type = ty } | GHole (loc, k, naming, None) -> let env = ltac_interp_name_env k0 lvar env in let ty = match tycon with - | Some ty -> ty + | Some ty -> EConstr.Unsafe.to_constr ty | None -> new_type_evar env evdref loc in - { uj_val = e_new_evar env evdref ~src:(loc,k) ~naming ty; uj_type = ty } + { uj_val = e_new_evar env evdref ~src:(loc,k) ~naming (EConstr.of_constr ty); uj_type = ty } | GHole (loc, k, _naming, Some arg) -> let env = ltac_interp_name_env k0 lvar env in let ty = match tycon with - | Some ty -> ty + | Some ty -> EConstr.Unsafe.to_constr ty | None -> new_type_evar env evdref loc in let ist = lvar.ltac_genargs in @@ -616,8 +637,8 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre (fun e ar -> pretype_type empty_valcon (push_rel_context e env) evdref lvar ar) ctxtv lar in - let lara = Array.map (fun a -> a.utj_val) larj in - let ftys = Array.map2 (fun e a -> it_mkProd_or_LetIn a e) ctxtv lara in + let lara = Array.map (fun a -> EConstr.of_constr a.utj_val) larj in + let ftys = Array.map2 (fun e a -> EConstr.it_mkProd_or_LetIn a e) ctxtv lara in let nbfix = Array.length lar in let names = Array.map (fun id -> Name id) names in let _ = @@ -626,7 +647,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre let fixi = match fixkind with | GFix (vn,i) -> i | GCoFix i -> i - in e_conv env.ExtraEnv.env evdref (EConstr.of_constr ftys.(fixi)) (EConstr.of_constr t) + in e_conv env.ExtraEnv.env evdref ftys.(fixi) t | None -> true in (* Note: bodies are not used by push_rec_types, so [||] is safe *) @@ -637,16 +658,17 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre (* we lift nbfix times the type in tycon, because of * the nbfix variables pushed to newenv *) let (ctxt,ty) = - decompose_prod_n_assum (Context.Rel.length ctxt) + decompose_prod_n_assum !evdref (Context.Rel.length ctxt) (lift nbfix ftys.(i)) in let nenv = push_rel_context ctxt newenv in - let j = pretype (mk_tycon (EConstr.of_constr ty)) nenv evdref lvar def in - { uj_val = it_mkLambda_or_LetIn j.uj_val ctxt; - uj_type = it_mkProd_or_LetIn j.uj_type ctxt }) + let j = pretype (mk_tycon ty) nenv evdref lvar def in + { uj_val = Termops.it_mkLambda_or_LetIn j.uj_val ctxt; + uj_type = Termops.it_mkProd_or_LetIn j.uj_type ctxt }) ctxtv vdef in - Typing.check_type_fixpoint loc env.ExtraEnv.env evdref names (Array.map EConstr.of_constr ftys) vdefj; - let ftys = Array.map (nf_evar !evdref) ftys in - let fdefs = Array.map (fun x -> nf_evar !evdref (j_val x)) vdefj in + Typing.check_type_fixpoint loc env.ExtraEnv.env evdref names ftys vdefj; + let nf c = nf_evar !evdref c in + let ftys = Array.map nf ftys in (** FIXME *) + let fdefs = Array.map (fun x -> nf (j_val x)) vdefj in let fixj = match fixkind with | GFix (vn,i) -> (* First, let's find the guard indexes. *) @@ -665,12 +687,13 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre let fixdecls = (names,ftys,fdefs) in let indexes = search_guard - loc env.ExtraEnv.env possible_indexes fixdecls + loc env.ExtraEnv.env possible_indexes (nf_fix !evdref fixdecls) in make_judge (mkFix ((indexes,i),fixdecls)) ftys.(i) | GCoFix i -> - let cofix = (i,(names,ftys,fdefs)) in - (try check_cofix env.ExtraEnv.env cofix + let fixdecls = (names,ftys,fdefs) in + let cofix = (i, fixdecls) in + (try check_cofix env.ExtraEnv.env (i, nf_fix !evdref fixdecls) with reraise -> let (e, info) = CErrors.push reraise in let info = Loc.add_loc info loc in @@ -691,24 +714,24 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre (* Bidirectional typechecking hint: parameters of a constructor are completely determined by a typing constraint *) - if Flags.is_program_mode () && length > 0 && isConstruct fj.uj_val then + if Flags.is_program_mode () && length > 0 && isConstruct !evdref (EConstr.of_constr fj.uj_val) then match tycon with | None -> [] | Some ty -> - let ((ind, i), u) = destConstruct fj.uj_val in + let ((ind, i), u) = destConstruct !evdref (EConstr.of_constr fj.uj_val) in let npars = inductive_nparams ind in if Int.equal npars 0 then [] else try - let IndType (indf, args) = find_rectype env.ExtraEnv.env !evdref (EConstr.of_constr ty) in + let IndType (indf, args) = find_rectype env.ExtraEnv.env !evdref ty in let ((ind',u'),pars) = dest_ind_family indf in - if eq_ind ind ind' then pars + if eq_ind ind ind' then List.map EConstr.of_constr pars else (* Let the usual code throw an error *) [] with Not_found -> [] else [] in let app_f = - match kind_of_term fj.uj_val with + match EConstr.kind !evdref (EConstr.of_constr fj.uj_val) with | Const (p, u) when Environ.is_projection p env.ExtraEnv.env -> let p = Projection.make p false in let pb = Environ.lookup_projection p env.ExtraEnv.env in @@ -724,7 +747,8 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre let argloc = loc_of_glob_constr c in let resj = evd_comb1 (Coercion.inh_app_fun resolve_tc env.ExtraEnv.env) evdref resj in let resty = whd_all env.ExtraEnv.env !evdref (EConstr.of_constr resj.uj_type) in - match kind_of_term resty with + let resty = EConstr.of_constr resty in + match EConstr.kind !evdref resty with | Prod (na,c1,c2) -> let tycon = Some c1 in let hj = pretype tycon env evdref lvar c in @@ -732,12 +756,12 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre match candargs with | [] -> [], j_val hj | arg :: args -> - if e_conv env.ExtraEnv.env evdref (EConstr.of_constr (j_val hj)) (EConstr.of_constr arg) then + if e_conv env.ExtraEnv.env evdref (j_val hj) arg then args, nf_evar !evdref (j_val hj) else [], j_val hj in let value, typ = app_f n (j_val resj) ujval, subst1 ujval c2 in - let j = { uj_val = value; uj_type = typ } in + let j = { uj_val = EConstr.Unsafe.to_constr value; uj_type = EConstr.Unsafe.to_constr typ } in apply_rec env (n+1) j candargs rest | _ -> @@ -748,16 +772,16 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre in let resj = apply_rec env 1 fj candargs args in let resj = - match evar_kind_of_term !evdref resj.uj_val with + match EConstr.kind !evdref (EConstr.of_constr resj.uj_val) with | App (f,args) -> - let f = whd_evar !evdref f in - if is_template_polymorphic env.ExtraEnv.env !evdref (EConstr.of_constr f) then + if is_template_polymorphic env.ExtraEnv.env !evdref f then (* Special case for inductive type applications that must be refreshed right away. *) - let sigma = !evdref in - let c = mkApp (f,Array.map (whd_evar sigma) args) in - let c = evd_comb1 (Evarsolve.refresh_universes (Some true) env.ExtraEnv.env) evdref (EConstr.of_constr c) in - let t = Retyping.get_type_of env.ExtraEnv.env !evdref (EConstr.of_constr c) in + let c = mkApp (f, args) in + let c = evd_comb1 (Evarsolve.refresh_universes (Some true) env.ExtraEnv.env) evdref c in + let c = EConstr.of_constr c in + let t = Retyping.get_type_of env.ExtraEnv.env !evdref c in + let t = EConstr.of_constr t in make_judge c (* use this for keeping evars: resj.uj_val *) t else resj | _ -> resj @@ -770,8 +794,8 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre match tycon with | None -> evd, tycon | Some ty -> - let evd, ty' = Coercion.inh_coerce_to_prod loc env.ExtraEnv.env evd (EConstr.of_constr ty) in - evd, Some (EConstr.Unsafe.to_constr ty')) + let evd, ty' = Coercion.inh_coerce_to_prod loc env.ExtraEnv.env evd ty in + evd, Some ty') evdref tycon in let (name',dom,rng) = evd_comb1 (split_tycon loc env.ExtraEnv.env) evdref tycon' in @@ -794,7 +818,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre let j' = match name with | Anonymous -> let j = pretype_type empty_valcon env evdref lvar c2 in - { j with utj_val = lift 1 j.utj_val } + { j with utj_val = EConstr.Unsafe.to_constr (lift 1 (EConstr.of_constr j.utj_val)) } | Name _ -> let var = LocalAssum (name, j.utj_val) in let env' = push_rel var env in @@ -825,11 +849,12 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre the substitution must also be applied on variables before they are looked up in the rel context. *) let var = LocalDef (name, j.uj_val, t) in + let t = EConstr.of_constr t in let tycon = lift_tycon 1 tycon in let j' = pretype tycon (push_rel var env) evdref lvar c2 in let name = ltac_interp_name lvar name in - { uj_val = mkLetIn (name, j.uj_val, t, j'.uj_val) ; - uj_type = subst1 j.uj_val j'.uj_type } + { uj_val = EConstr.Unsafe.to_constr (mkLetIn (name, EConstr.of_constr j.uj_val, t, EConstr.of_constr j'.uj_val)) ; + uj_type = EConstr.Unsafe.to_constr (subst1 (EConstr.of_constr j.uj_val) (EConstr.of_constr j'.uj_type)) } | GLetTuple (loc,nal,(na,po),c,d) -> let cj = pretype empty_tycon env evdref lvar c in @@ -839,6 +864,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre let cloc = loc_of_glob_constr c in error_case_not_inductive ~loc:cloc env.ExtraEnv.env !evdref cj in + let realargs = List.map EConstr.of_constr realargs in let cstrs = get_constructors env.ExtraEnv.env indf in if not (Int.equal (Array.length cstrs) 1) then user_err ~loc (str "Destructing let is only for inductive types" ++ @@ -855,8 +881,9 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre let rec aux n k names l = match names, l with | na :: names, (LocalAssum (_,t) :: l) -> + let t = EConstr.of_constr t in let proj = Projection.make ps.(cs.cs_nargs - k) true in - LocalDef (na, lift (cs.cs_nargs - n) (mkProj (proj, cj.uj_val)), t) + local_def (na, lift (cs.cs_nargs - n) (mkProj (proj, EConstr.of_constr cj.uj_val)), t) :: aux (n+1) (k + 1) names l | na :: names, (decl :: l) -> set_name na decl :: aux (n+1) k names l @@ -870,7 +897,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre let fsign = List.map2 set_name nal fsign in let f = it_mkLambda_or_LetIn f fsign in let ci = make_case_info env.ExtraEnv.env (fst ind) LetStyle in - mkCase (ci, p, cj.uj_val,[|f|]) + mkCase (ci, p, EConstr.of_constr cj.uj_val,[|f|]) else it_mkLambda_or_LetIn f fsign in let env_f = push_rel_context fsign env in @@ -887,28 +914,28 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre | Some p -> let env_p = push_rel_context psign env in let pj = pretype_type empty_valcon env_p evdref lvar p in - let ccl = nf_evar !evdref pj.utj_val in + let ccl = nf_evar !evdref (EConstr.of_constr pj.utj_val) in let psign = make_arity_signature env.ExtraEnv.env true indf in (* with names *) let p = it_mkLambda_or_LetIn ccl psign in let inst = - (Array.to_list cs.cs_concl_realargs) - @[build_dependent_constructor cs] in + (Array.map_to_list EConstr.of_constr cs.cs_concl_realargs) + @[EConstr.of_constr (build_dependent_constructor cs)] in let lp = lift cs.cs_nargs p in - let fty = hnf_lam_applist env.ExtraEnv.env !evdref (EConstr.of_constr lp) (List.map EConstr.of_constr inst) in + let fty = hnf_lam_applist env.ExtraEnv.env !evdref lp inst in let fj = pretype (mk_tycon (EConstr.of_constr fty)) env_f evdref lvar d in let v = let ind,_ = dest_ind_family indf in - Typing.check_allowed_sort env.ExtraEnv.env !evdref ind (EConstr.of_constr cj.uj_val) (EConstr.of_constr p); - obj ind p cj.uj_val fj.uj_val + Typing.check_allowed_sort env.ExtraEnv.env !evdref ind (EConstr.of_constr cj.uj_val) p; + obj ind p cj.uj_val (EConstr.of_constr fj.uj_val) in - { uj_val = v; uj_type = substl (realargs@[cj.uj_val]) ccl } + { uj_val = EConstr.Unsafe.to_constr v; uj_type = EConstr.Unsafe.to_constr (substl (realargs@[EConstr.of_constr cj.uj_val]) ccl) } | None -> let tycon = lift_tycon cs.cs_nargs tycon in let fj = pretype tycon env_f evdref lvar d in - let ccl = nf_evar !evdref fj.uj_type in + let ccl = nf_evar !evdref (EConstr.of_constr fj.uj_type) in let ccl = - if noccur_between 1 cs.cs_nargs ccl then + if noccur_between !evdref 1 cs.cs_nargs ccl then lift (- cs.cs_nargs) ccl else error_cant_find_case_type ~loc env.ExtraEnv.env !evdref @@ -917,9 +944,9 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre let p = it_mkLambda_or_LetIn (lift (nar+1) ccl) psign in let v = let ind,_ = dest_ind_family indf in - Typing.check_allowed_sort env.ExtraEnv.env !evdref ind (EConstr.of_constr cj.uj_val) (EConstr.of_constr p); - obj ind p cj.uj_val fj.uj_val - in { uj_val = v; uj_type = ccl }) + Typing.check_allowed_sort env.ExtraEnv.env !evdref ind (EConstr.of_constr cj.uj_val) p; + obj ind p cj.uj_val (EConstr.of_constr fj.uj_val) + in { uj_val = EConstr.Unsafe.to_constr v; uj_type = EConstr.Unsafe.to_constr ccl }) | GIf (loc,c,(na,po),b1,b2) -> let cj = pretype empty_tycon env evdref lvar c in @@ -946,16 +973,16 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre | Some p -> let env_p = push_rel_context psign env in let pj = pretype_type empty_valcon env_p evdref lvar p in - let ccl = nf_evar !evdref pj.utj_val in + let ccl = nf_evar !evdref (EConstr.of_constr pj.utj_val) in let pred = it_mkLambda_or_LetIn ccl psign in - let typ = lift (- nar) (beta_applist !evdref (EConstr.of_constr pred,[EConstr.of_constr cj.uj_val])) in + let typ = lift (- nar) (EConstr.of_constr (beta_applist !evdref (pred,[EConstr.of_constr cj.uj_val]))) in pred, typ | None -> let p = match tycon with | Some ty -> ty | None -> let env = ltac_interp_name_env k0 lvar env in - new_type_evar env evdref loc + EConstr.of_constr (new_type_evar env evdref loc) in it_mkLambda_or_LetIn (lift (nar+1) p) psign, p in let pred = nf_evar !evdref pred in @@ -963,7 +990,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre let f cs b = let n = Context.Rel.length cs.cs_args in let pi = lift n pred in (* liftn n 2 pred ? *) - let pi = beta_applist !evdref (EConstr.of_constr pi, [EConstr.of_constr (build_dependent_constructor cs)]) in + let pi = beta_applist !evdref (pi, [EConstr.of_constr (build_dependent_constructor cs)]) in let csgn = if not !allow_anonymous_refs then List.map (set_name Anonymous) cs.cs_args @@ -974,17 +1001,17 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre in let env_c = push_rel_context csgn env in let bj = pretype (mk_tycon (EConstr.of_constr pi)) env_c evdref lvar b in - it_mkLambda_or_LetIn bj.uj_val cs.cs_args in + it_mkLambda_or_LetIn (EConstr.of_constr bj.uj_val) cs.cs_args in let b1 = f cstrs.(0) b1 in let b2 = f cstrs.(1) b2 in let v = let ind,_ = dest_ind_family indf in let ci = make_case_info env.ExtraEnv.env (fst ind) IfStyle in let pred = nf_evar !evdref pred in - Typing.check_allowed_sort env.ExtraEnv.env !evdref ind (EConstr.of_constr cj.uj_val) (EConstr.of_constr pred); - mkCase (ci, pred, cj.uj_val, [|b1;b2|]) + Typing.check_allowed_sort env.ExtraEnv.env !evdref ind (EConstr.of_constr cj.uj_val) pred; + mkCase (ci, pred, EConstr.of_constr cj.uj_val, [|b1;b2|]) in - let cj = { uj_val = v; uj_type = p } in + let cj = { uj_val = EConstr.Unsafe.to_constr v; uj_type = EConstr.Unsafe.to_constr p } in inh_conv_coerce_to_tycon loc env evdref cj tycon | GCases (loc,sty,po,tml,eqns) -> @@ -1004,53 +1031,56 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre let tval = evd_comb1 (Evarsolve.refresh_universes ~onlyalg:true ~status:Evd.univ_flexible (Some false) env.ExtraEnv.env) evdref (EConstr.of_constr tj.utj_val) in + let tval = EConstr.of_constr tval in let tval = nf_evar !evdref tval in let cj, tval = match k with | VMcast -> let cj = pretype empty_tycon env evdref lvar c in - let cty = nf_evar !evdref cj.uj_type and tval = nf_evar !evdref tval in - if not (occur_existential !evdref (EConstr.of_constr cty) || occur_existential !evdref (EConstr.of_constr tval)) then + let cty = nf_evar !evdref (EConstr.of_constr cj.uj_type) and tval = nf_evar !evdref tval in + if not (occur_existential !evdref cty || occur_existential !evdref tval) then let (evd,b) = Reductionops.vm_infer_conv env.ExtraEnv.env !evdref cty tval in if b then (evdref := evd; cj, tval) else - error_actual_type ~loc env.ExtraEnv.env !evdref cj tval - (ConversionFailed (env.ExtraEnv.env,EConstr.of_constr cty,EConstr.of_constr tval)) + error_actual_type ~loc env.ExtraEnv.env !evdref cj (EConstr.Unsafe.to_constr tval) + (ConversionFailed (env.ExtraEnv.env,cty,tval)) else user_err ~loc (str "Cannot check cast with vm: " ++ str "unresolved arguments remain.") | NATIVEcast -> let cj = pretype empty_tycon env evdref lvar c in - let cty = nf_evar !evdref cj.uj_type and tval = nf_evar !evdref tval in + let cty = nf_evar !evdref (EConstr.of_constr cj.uj_type) and tval = nf_evar !evdref tval in begin - let (evd,b) = Nativenorm.native_infer_conv env.ExtraEnv.env !evdref (EConstr.of_constr cty) (EConstr.of_constr tval) in + let (evd,b) = Nativenorm.native_infer_conv env.ExtraEnv.env !evdref cty tval in if b then (evdref := evd; cj, tval) else - error_actual_type ~loc env.ExtraEnv.env !evdref cj tval - (ConversionFailed (env.ExtraEnv.env,EConstr.of_constr cty,EConstr.of_constr tval)) + error_actual_type ~loc env.ExtraEnv.env !evdref cj (EConstr.Unsafe.to_constr tval) + (ConversionFailed (env.ExtraEnv.env,cty,tval)) end | _ -> - pretype (mk_tycon (EConstr.of_constr tval)) env evdref lvar c, tval + pretype (mk_tycon tval) env evdref lvar c, tval in - let v = mkCast (cj.uj_val, k, tval) in - { uj_val = v; uj_type = tval } + let v = mkCast (EConstr.of_constr cj.uj_val, k, tval) in + { uj_val = EConstr.Unsafe.to_constr v; uj_type = EConstr.Unsafe.to_constr tval } in inh_conv_coerce_to_tycon loc env evdref cj tycon and pretype_instance k0 resolve_tc env evdref lvar loc hyps evk update = let f decl (subst,update) = let id = NamedDecl.get_id decl in - let t = replace_vars subst (NamedDecl.get_type decl) in + let t = replace_vars subst (EConstr.of_constr (NamedDecl.get_type decl)) in let c, update = try let c = List.assoc id update in - let c = pretype k0 resolve_tc (mk_tycon (EConstr.of_constr t)) env evdref lvar c in - c.uj_val, List.remove_assoc id update + let c = pretype k0 resolve_tc (mk_tycon t) env evdref lvar c in + EConstr.of_constr c.uj_val, List.remove_assoc id update with Not_found -> try let (n,_,t') = lookup_rel_id id (rel_context env) in - if is_conv env.ExtraEnv.env !evdref (EConstr.of_constr t) (EConstr.of_constr t') then mkRel n, update else raise Not_found + let t' = EConstr.of_constr t' in + if is_conv env.ExtraEnv.env !evdref t t' then mkRel n, update else raise Not_found with Not_found -> try let t' = env |> lookup_named id |> NamedDecl.get_type in - if is_conv env.ExtraEnv.env !evdref (EConstr.of_constr t) (EConstr.of_constr t') then mkVar id, update else raise Not_found + let t' = EConstr.of_constr t' in + if is_conv env.ExtraEnv.env !evdref t t' then mkVar id, update else raise Not_found with Not_found -> user_err ~loc (str "Cannot interpret " ++ pr_existential_key !evdref evk ++ @@ -1063,18 +1093,23 @@ and pretype_instance k0 resolve_tc env evdref lvar loc hyps evk update = (* [pretype_type valcon env evdref lvar c] coerces [c] into a type *) and pretype_type k0 resolve_tc valcon (env : ExtraEnv.t) evdref lvar = function | GHole (loc, knd, naming, None) -> + let rec is_Type c = match EConstr.kind !evdref c with + | Sort (Type _) -> true + | Cast (c, _, _) -> is_Type c + | _ -> false + in (match valcon with | Some v -> let s = let sigma = !evdref in - let t = Retyping.get_type_of env.ExtraEnv.env sigma (EConstr.of_constr v) in + let t = Retyping.get_type_of env.ExtraEnv.env sigma v in match EConstr.kind sigma (EConstr.of_constr (whd_all env.ExtraEnv.env sigma (EConstr.of_constr t))) with | Sort s -> s - | Evar ev when is_Type (existential_type sigma (fst ev, Array.map EConstr.Unsafe.to_constr (snd ev))) -> + | Evar ev when is_Type (existential_type sigma ev) -> evd_comb1 (define_evar_as_sort env.ExtraEnv.env) evdref ev | _ -> anomaly (Pp.str "Found a type constraint which is not a type") in - { utj_val = v; + { utj_val = EConstr.Unsafe.to_constr v; utj_type = s } | None -> let env = ltac_interp_name_env k0 lvar env in @@ -1088,10 +1123,10 @@ and pretype_type k0 resolve_tc valcon (env : ExtraEnv.t) evdref lvar = function match valcon with | None -> tj | Some v -> - if e_cumul env.ExtraEnv.env evdref (EConstr.of_constr v) (EConstr.of_constr tj.utj_val) then tj + if e_cumul env.ExtraEnv.env evdref v (EConstr.of_constr tj.utj_val) then tj else error_unexpected_type - ~loc:(loc_of_glob_constr c) env.ExtraEnv.env !evdref tj.utj_val v + ~loc:(loc_of_glob_constr c) env.ExtraEnv.env !evdref tj.utj_val (EConstr.Unsafe.to_constr v) let ise_pretype_gen flags env sigma lvar kind c = let env = make_env env in @@ -1101,11 +1136,11 @@ let ise_pretype_gen flags env sigma lvar kind c = | WithoutTypeConstraint -> (pretype k0 flags.use_typeclasses empty_tycon env evdref lvar c).uj_val | OfType exptyp -> - (pretype k0 flags.use_typeclasses (mk_tycon (EConstr.of_constr exptyp)) env evdref lvar c).uj_val + (pretype k0 flags.use_typeclasses (mk_tycon exptyp) env evdref lvar c).uj_val | IsType -> (pretype_type k0 flags.use_typeclasses empty_valcon env evdref lvar c).utj_val in - process_inference_flags flags env.ExtraEnv.env sigma (!evdref,c') + process_inference_flags flags env.ExtraEnv.env sigma (!evdref,EConstr.of_constr c') let default_inference_flags fail = { use_typeclasses = true; @@ -1131,17 +1166,17 @@ let empty_lvar : ltac_var_map = { ltac_genargs = Id.Map.empty; } -let on_judgment f j = - let c = mkCast(j.uj_val,DEFAULTcast, j.uj_type) in - let (c,_,t) = destCast (f c) in - {uj_val = c; uj_type = t} +let on_judgment sigma f j = + let c = mkCast(EConstr.of_constr j.uj_val,DEFAULTcast, EConstr.of_constr j.uj_type) in + let (c,_,t) = destCast sigma (f c) in + {uj_val = EConstr.Unsafe.to_constr c; uj_type = EConstr.Unsafe.to_constr t} let understand_judgment env sigma c = let env = make_env env in let evdref = ref sigma in let k0 = Context.Rel.length (rel_context env) in let j = pretype k0 true empty_tycon env evdref empty_lvar c in - let j = on_judgment (fun c -> + let j = on_judgment sigma (fun c -> let evd, c = process_inference_flags all_and_fail_flags env.ExtraEnv.env sigma (!evdref,c) in evdref := evd; c) j in j, Evd.evar_universe_context !evdref @@ -1150,14 +1185,14 @@ let understand_judgment_tcc env evdref c = let env = make_env env in let k0 = Context.Rel.length (rel_context env) in let j = pretype k0 true empty_tycon env evdref empty_lvar c in - on_judgment (fun c -> + on_judgment !evdref (fun c -> let (evd,c) = process_inference_flags all_no_fail_flags env.ExtraEnv.env Evd.empty (!evdref,c) in evdref := evd; c) j let ise_pretype_gen_ctx flags env sigma lvar kind c = let evd, c = ise_pretype_gen flags env sigma lvar kind c in let evd, f = Evarutil.nf_evars_and_universes evd in - f c, Evd.evar_universe_context evd + f (EConstr.Unsafe.to_constr c), Evd.evar_universe_context evd (** Entry points of the high-level type synthesis algorithm *) @@ -1168,15 +1203,17 @@ let understand ise_pretype_gen_ctx flags env sigma empty_lvar expected_type c let understand_tcc ?(flags=all_no_fail_flags) env sigma ?(expected_type=WithoutTypeConstraint) c = - ise_pretype_gen flags env sigma empty_lvar expected_type c + let (sigma, c) = ise_pretype_gen flags env sigma empty_lvar expected_type c in + (sigma, EConstr.Unsafe.to_constr c) let understand_tcc_evars ?(flags=all_no_fail_flags) env evdref ?(expected_type=WithoutTypeConstraint) c = let sigma, c = ise_pretype_gen flags env !evdref empty_lvar expected_type c in evdref := sigma; - c + EConstr.Unsafe.to_constr c let understand_ltac flags env sigma lvar kind c = - ise_pretype_gen flags env sigma lvar kind c + let (sigma, c) = ise_pretype_gen flags env sigma lvar kind c in + (sigma, EConstr.Unsafe.to_constr c) let constr_flags = { use_typeclasses = true; diff --git a/pretyping/pretyping.mli b/pretyping/pretyping.mli index e09648ec3..603b9f9ea 100644 --- a/pretyping/pretyping.mli +++ b/pretyping/pretyping.mli @@ -25,7 +25,7 @@ open Misctypes val search_guard : Loc.t -> env -> int list list -> rec_declaration -> int array -type typing_constraint = OfType of types | IsType | WithoutTypeConstraint +type typing_constraint = OfType of EConstr.types | IsType | WithoutTypeConstraint type var_map = Pattern.constr_under_binders Id.Map.t type uconstr_var_map = Glob_term.closed_glob_constr Id.Map.t @@ -47,7 +47,7 @@ val empty_lvar : ltac_var_map type glob_constr_ltac_closure = ltac_var_map * glob_constr type pure_open_constr = evar_map * constr -type inference_hook = env -> evar_map -> evar -> evar_map * constr +type inference_hook = env -> evar_map -> evar -> evar_map * EConstr.constr type inference_flags = { use_typeclasses : bool; @@ -139,7 +139,7 @@ val check_evars_are_solved : (** [check_evars env initial_sigma extended_sigma c] fails if some new unresolved evar remains in [c] *) -val check_evars : env -> evar_map -> evar_map -> constr -> unit +val check_evars : env -> evar_map -> evar_map -> EConstr.constr -> unit (**/**) (** Internal of Pretyping... *) @@ -153,7 +153,7 @@ val pretype_type : val ise_pretype_gen : inference_flags -> env -> evar_map -> - ltac_var_map -> typing_constraint -> glob_constr -> evar_map * constr + ltac_var_map -> typing_constraint -> glob_constr -> evar_map * EConstr.constr (**/**) diff --git a/pretyping/reductionops.ml b/pretyping/reductionops.ml index 69d47e8e6..0b97cd253 100644 --- a/pretyping/reductionops.ml +++ b/pretyping/reductionops.ml @@ -1317,6 +1317,8 @@ let sigma_univ_state = let infer_conv_gen conv_fun ?(catch_incon=true) ?(pb=Reduction.CUMUL) ?(ts=full_transparent_state) env sigma x y = + let x = EConstr.Unsafe.to_constr x in + let y = EConstr.Unsafe.to_constr y in try let fold cstr sigma = try Some (Evd.add_universe_constraints sigma cstr) diff --git a/pretyping/reductionops.mli b/pretyping/reductionops.mli index 911dab0b6..5e6a40786 100644 --- a/pretyping/reductionops.mli +++ b/pretyping/reductionops.mli @@ -264,12 +264,12 @@ val check_conv : ?pb:conv_pb -> ?ts:transparent_state -> env -> evar_map -> ECo otherwise returns false in that case. *) val infer_conv : ?catch_incon:bool -> ?pb:conv_pb -> ?ts:transparent_state -> - env -> evar_map -> constr -> constr -> evar_map * bool + env -> evar_map -> EConstr.constr -> EConstr.constr -> evar_map * bool (** Conversion with inference of universe constraints *) -val set_vm_infer_conv : (?pb:conv_pb -> env -> evar_map -> constr -> constr -> +val set_vm_infer_conv : (?pb:conv_pb -> env -> evar_map -> EConstr.constr -> EConstr.constr -> evar_map * bool) -> unit -val vm_infer_conv : ?pb:conv_pb -> env -> evar_map -> constr -> constr -> +val vm_infer_conv : ?pb:conv_pb -> env -> evar_map -> EConstr.constr -> EConstr.constr -> evar_map * bool @@ -278,7 +278,7 @@ conversion function. Used to pretype vm and native casts. *) val infer_conv_gen : (conv_pb -> l2r:bool -> evar_map -> transparent_state -> (constr, evar_map) Reduction.generic_conversion_function) -> ?catch_incon:bool -> ?pb:conv_pb -> ?ts:transparent_state -> env -> - evar_map -> constr -> constr -> evar_map * bool + evar_map -> EConstr.constr -> EConstr.constr -> evar_map * bool (** {6 Special-Purpose Reduction Functions } *) diff --git a/pretyping/tacred.ml b/pretyping/tacred.ml index b729f3b9b..5b8eaa50b 100644 --- a/pretyping/tacred.ml +++ b/pretyping/tacred.ml @@ -413,7 +413,7 @@ let substl_with_function subst sigma constr = match v.(i-k-1) with | (fx, Some (min, ref)) -> let sigma = Sigma.Unsafe.of_evar_map !evd in - let Sigma (evk, sigma, _) = Evarutil.new_pure_evar venv sigma dummy in + let Sigma (evk, sigma, _) = Evarutil.new_pure_evar venv sigma (EConstr.of_constr dummy) in let sigma = Sigma.to_evar_map sigma in evd := sigma; minargs := Evar.Map.add evk min !minargs; diff --git a/pretyping/unification.ml b/pretyping/unification.ml index b568dd044..70aa0be6b 100644 --- a/pretyping/unification.ml +++ b/pretyping/unification.ml @@ -167,6 +167,7 @@ let pose_all_metas_as_evars env evd t = let {rebus=ty;freemetas=mvs} = Evd.meta_ftype evd mv in let ty = if Evd.Metaset.is_empty mvs then ty else aux ty in let src = Evd.evar_source_of_meta mv !evdref in + let ty = EConstr.of_constr ty in let ev = Evarutil.e_new_evar env evdref ~src ty in evdref := meta_assign mv (ev,(Conv,TypeNotProcessed)) !evdref; ev) @@ -608,7 +609,7 @@ let check_compatibility env pbty flags (sigma,metasubst,evarsubst) tyM tyN = | None -> sigma | Some n -> if is_ground_term sigma (EConstr.of_constr m) && is_ground_term sigma (EConstr.of_constr n) then - let sigma, b = infer_conv ~pb:pbty ~ts:flags.modulo_delta_types env sigma m n in + let sigma, b = infer_conv ~pb:pbty ~ts:flags.modulo_delta_types env sigma (EConstr.of_constr m) (EConstr.of_constr n) in if b then sigma else error_cannot_unify env sigma (m,n) else sigma @@ -961,10 +962,11 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb (* Renounce, maybe metas/evars prevents typing *) sigma else sigma in + let m1 = EConstr.of_constr m1 and n1 = EConstr.of_constr n1 in let sigma, b = infer_conv ~pb ~ts:convflags curenv sigma m1 n1 in if b then Some (sigma, metasubst, evarsubst) else - if is_ground_term sigma (EConstr.of_constr m1) && is_ground_term sigma (EConstr.of_constr n1) then + if is_ground_term sigma m1 && is_ground_term sigma n1 then error_cannot_unify curenv sigma (cM,cN) else None in @@ -1071,7 +1073,7 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb None else let sigma, b = match flags.modulo_conv_on_closed_terms with - | Some convflags -> infer_conv ~pb:cv_pb ~ts:convflags env sigma m n + | Some convflags -> infer_conv ~pb:cv_pb ~ts:convflags env sigma (EConstr.of_constr m) (EConstr.of_constr n) | _ -> constr_cmp cv_pb sigma flags m n in if b then Some sigma else if (match flags.modulo_conv_on_closed_terms, flags.modulo_delta with @@ -1210,7 +1212,7 @@ let applyHead env (type r) (evd : r Sigma.t) n c = else match kind_of_term (whd_all env (Sigma.to_evar_map evd) (EConstr.of_constr cty)) with | Prod (_,c1,c2) -> - let Sigma (evar, evd', q) = Evarutil.new_evar env evd ~src:(Loc.ghost,Evar_kinds.GoalEvar) c1 in + let Sigma (evar, evd', q) = Evarutil.new_evar env evd ~src:(Loc.ghost,Evar_kinds.GoalEvar) (EConstr.of_constr c1) in apprec (n-1) (mkApp(c,[|evar|])) (subst1 evar c2) (p +> q) evd' | _ -> error "Apply_Head_Then" in @@ -1570,7 +1572,7 @@ let make_pattern_test from_prefix_of_ind is_correct_type env sigma (pending,c) = let merge_fun c1 c2 = match c1, c2 with | Some (evd,c1,x), Some (_,c2,_) -> - let (evd,b) = infer_conv ~pb:CONV env evd (EConstr.Unsafe.to_constr c1) (EConstr.Unsafe.to_constr c2) in + let (evd,b) = infer_conv ~pb:CONV env evd c1 c2 in if b then Some (evd, c1, x) else raise (NotUnifiable None) | Some _, None -> c1 | None, Some _ -> c2 @@ -1883,7 +1885,7 @@ let secondOrderAbstraction env evd flags typ (p, oplist) = let (evd',cllist) = w_unify_to_subterm_list env evd flags p oplist typ in let typp = Typing.meta_type evd' p in let evd',(pred,predtyp) = abstract_list_all env evd' typp typ cllist in - let evd', b = infer_conv ~pb:CUMUL env evd' predtyp typp in + let evd', b = infer_conv ~pb:CUMUL env evd' (EConstr.of_constr predtyp) (EConstr.of_constr typp) in if not b then error_wrong_abstraction_type env evd' (Evd.meta_name evd p) pred typp predtyp; @@ -1900,7 +1902,7 @@ let secondOrderAbstraction env evd flags typ (p, oplist) = let secondOrderDependentAbstraction env evd flags typ (p, oplist) = let typp = Typing.meta_type evd p in - let evd, pred = abstract_list_all_with_dependencies env evd typp typ (List.map EConstr.of_constr oplist) in + let evd, pred = abstract_list_all_with_dependencies env evd (EConstr.of_constr typp) typ (List.map EConstr.of_constr oplist) in w_merge env false flags.merge_unify_flags (evd,[p,pred,(Conv,TypeProcessed)],[]) diff --git a/proofs/clenv.ml b/proofs/clenv.ml index 0515e4198..68aeaef5e 100644 --- a/proofs/clenv.ml +++ b/proofs/clenv.ml @@ -337,7 +337,7 @@ let clenv_pose_metas_as_evars clenv dep_mvs = let src = evar_source_of_meta mv clenv.evd in let src = adjust_meta_source clenv.evd mv src in let evd = Sigma.Unsafe.of_evar_map clenv.evd in - let Sigma (evar, evd, _) = new_evar (cl_env clenv) evd ~src ty in + let Sigma (evar, evd, _) = new_evar (cl_env clenv) evd ~src (EConstr.of_constr ty) in let evd = Sigma.to_evar_map evd in let clenv = clenv_assign mv evar {clenv with evd=evd} in fold clenv mvs in @@ -610,7 +610,7 @@ let make_evar_clause env sigma ?len t = | Prod (na, t1, t2) -> let store = Typeclasses.set_resolvable Evd.Store.empty false in let sigma = Sigma.Unsafe.of_evar_map sigma in - let Sigma (ev, sigma, _) = new_evar ~store env sigma t1 in + let Sigma (ev, sigma, _) = new_evar ~store env sigma (EConstr.of_constr t1) in let sigma = Sigma.to_evar_map sigma in let dep = not (EConstr.Vars.noccurn sigma 1 (EConstr.of_constr t2)) in let hole = { diff --git a/proofs/evar_refiner.ml b/proofs/evar_refiner.ml index af5fa921f..fd6528a1e 100644 --- a/proofs/evar_refiner.ml +++ b/proofs/evar_refiner.ml @@ -51,7 +51,7 @@ let w_refine (evk,evi) (ltac_var,rawc) sigma = Pretyping.fail_evar = false; Pretyping.expand_evars = true } in try Pretyping.understand_ltac flags - env sigma ltac_var (Pretyping.OfType evi.evar_concl) rawc + env sigma ltac_var (Pretyping.OfType (EConstr.of_constr evi.evar_concl)) rawc with e when CErrors.noncritical e -> let loc = Glob_ops.loc_of_glob_constr rawc in user_err ~loc diff --git a/proofs/logic.ml b/proofs/logic.ml index 93b31ced1..093d949d5 100644 --- a/proofs/logic.ml +++ b/proofs/logic.ml @@ -314,7 +314,7 @@ let check_meta_variables c = let check_conv_leq_goal env sigma arg ty conclty = if !check then - let evm, b = Reductionops.infer_conv env sigma ty conclty in + let evm, b = Reductionops.infer_conv env sigma (EConstr.of_constr ty) (EConstr.of_constr conclty) in if b then evm else raise (RefinerError (BadType (arg,ty,conclty))) else sigma diff --git a/proofs/pfedit.ml b/proofs/pfedit.ml index 67e216745..af910810b 100644 --- a/proofs/pfedit.ml +++ b/proofs/pfedit.ml @@ -236,6 +236,6 @@ let solve_by_implicit_tactic env sigma evk = let (ans, _, ctx) = build_by_tactic env (Evd.evar_universe_context sigma) c tac in let sigma = Evd.set_universe_context sigma ctx in - sigma, ans + sigma, EConstr.of_constr ans with e when Logic.catchable_exception e -> raise Exit) | _ -> raise Exit diff --git a/proofs/pfedit.mli b/proofs/pfedit.mli index 7458109fa..807a554df 100644 --- a/proofs/pfedit.mli +++ b/proofs/pfedit.mli @@ -190,4 +190,4 @@ val declare_implicit_tactic : unit Proofview.tactic -> unit val clear_implicit_tactic : unit -> unit (* Raise Exit if cannot solve *) -val solve_by_implicit_tactic : env -> Evd.evar_map -> Evd.evar -> Evd.evar_map * constr +val solve_by_implicit_tactic : env -> Evd.evar_map -> Evd.evar -> Evd.evar_map * EConstr.constr diff --git a/tactics/class_tactics.ml b/tactics/class_tactics.ml index ff7dbfa91..bc1d0ed6b 100644 --- a/tactics/class_tactics.ml +++ b/tactics/class_tactics.ml @@ -1458,7 +1458,7 @@ let _ = let resolve_one_typeclass env ?(sigma=Evd.empty) gl unique = let nc, gl, subst, _, _ = Evarutil.push_rel_context_to_named_context env gl in let (gl,t,sigma) = - Goal.V82.mk_goal sigma nc gl Store.empty in + Goal.V82.mk_goal sigma nc (EConstr.Unsafe.to_constr gl) Store.empty in let gls = { it = gl ; sigma = sigma; } in let hints = searchtable_map typeclasses_db in let st = Hint_db.transparent_state hints in @@ -1481,7 +1481,7 @@ let resolve_one_typeclass env ?(sigma=Evd.empty) gl unique = let _ = Hook.set Typeclasses.solve_one_instance_hook - (fun x y z w -> resolve_one_typeclass x ~sigma:y (EConstr.Unsafe.to_constr z) w) + (fun x y z w -> resolve_one_typeclass x ~sigma:y z w) (** Take the head of the arity of a constr. Used in the partial application tactic. *) diff --git a/tactics/equality.ml b/tactics/equality.ml index 9679ac402..be175937b 100644 --- a/tactics/equality.ml +++ b/tactics/equality.ml @@ -1176,7 +1176,7 @@ let sig_clausal_form env sigma sort_of_ty siglen ty dflt = let (a,p_i_minus_1) = match whd_beta_stack !evdref (EConstr.of_constr p_i) with | (_sigS,[a;p]) -> (EConstr.Unsafe.to_constr a, EConstr.Unsafe.to_constr p) | _ -> anomaly ~label:"sig_clausal_form" (Pp.str "should be a sigma type") in - let ev = Evarutil.e_new_evar env evdref a in + let ev = Evarutil.e_new_evar env evdref (EConstr.of_constr a) in let rty = beta_applist sigma (EConstr.of_constr p_i_minus_1,[EConstr.of_constr ev]) in let tuple_tail = sigrec_clausal_form (siglen-1) rty in match diff --git a/tactics/hints.ml b/tactics/hints.ml index b2aa02191..e8225df2d 100644 --- a/tactics/hints.ml +++ b/tactics/hints.ml @@ -1213,7 +1213,7 @@ let prepare_hint check (poly,local) env init (sigma,c) = subst := (evar,mkVar id)::!subst; mkNamedLambda id t (iter (replace_term sigma (EConstr.of_constr evar) (EConstr.mkVar id) (EConstr.of_constr c))) in let c' = iter c in - if check then Pretyping.check_evars (Global.env()) Evd.empty sigma c'; + if check then Pretyping.check_evars (Global.env()) Evd.empty sigma (EConstr.of_constr c'); let diff = Univ.ContextSet.diff (Evd.universe_context_set sigma) (Evd.universe_context_set init) in if poly then IsConstr (c', diff) else if local then IsConstr (c', diff) diff --git a/tactics/tactics.ml b/tactics/tactics.ml index a6bc805bd..3bb285aa8 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -176,7 +176,7 @@ let unsafe_intro env store decl b = let inst = List.map (NamedDecl.get_id %> mkVar) (named_context env) in let ninst = mkRel 1 :: inst in let nb = subst1 (mkVar (NamedDecl.get_id decl)) b in - let Sigma (ev, sigma, p) = new_evar_instance nctx sigma nb ~principal:true ~store ninst in + let Sigma (ev, sigma, p) = new_evar_instance nctx sigma (EConstr.of_constr nb) ~principal:true ~store ninst in Sigma (mkNamedLambda_or_LetIn decl ev, sigma, p) end } @@ -206,12 +206,13 @@ let convert_concl ?(check=true) ty k = let env = Proofview.Goal.env gl in let store = Proofview.Goal.extra gl in let conclty = Proofview.Goal.raw_concl gl in + let ty = EConstr.of_constr ty in Refine.refine ~unsafe:true { run = begin fun sigma -> let Sigma ((), sigma, p) = if check then begin let sigma = Sigma.to_evar_map sigma in - ignore (Typing.unsafe_type_of env sigma (EConstr.of_constr ty)); - let sigma,b = Reductionops.infer_conv env sigma ty conclty in + ignore (Typing.unsafe_type_of env sigma ty); + let sigma,b = Reductionops.infer_conv env sigma ty (EConstr.of_constr conclty) in if not b then error "Not convertible."; Sigma.Unsafe.of_pair ((), sigma) end else Sigma.here () sigma in @@ -230,7 +231,7 @@ let convert_hyp ?(check=true) d = let sign = convert_hyp check (named_context_val env) sigma d in let env = reset_with_named_context sign env in Refine.refine ~unsafe:true { run = begin fun sigma -> - Evarutil.new_evar env sigma ~principal:true ~store ty + Evarutil.new_evar env sigma ~principal:true ~store (EConstr.of_constr ty) end } end } @@ -248,8 +249,8 @@ let convert_gen pb x y = Tacticals.New.tclFAIL 0 (str "Not convertible") end } -let convert x y = convert_gen Reduction.CONV x y -let convert_leq x y = convert_gen Reduction.CUMUL x y +let convert x y = convert_gen Reduction.CONV (EConstr.of_constr x) (EConstr.of_constr y) +let convert_leq x y = convert_gen Reduction.CUMUL (EConstr.of_constr x) (EConstr.of_constr y) let clear_dependency_msg env sigma id = function | Evarutil.OccurHypInSimpleClause None -> @@ -300,7 +301,7 @@ let clear_gen fail = function in let env = reset_with_named_context hyps env in let tac = Refine.refine ~unsafe:true { run = fun sigma -> - Evarutil.new_evar env sigma ~principal:true concl + Evarutil.new_evar env sigma ~principal:true (EConstr.of_constr concl) } in Sigma.Unsafe.of_pair (tac, !evdref) end } @@ -330,7 +331,7 @@ let move_hyp id dest = let sign' = move_hyp_in_named_context sigma id dest sign in let env = reset_with_named_context sign' env in Refine.refine ~unsafe:true { run = begin fun sigma -> - Evarutil.new_evar env sigma ~principal:true ~store ty + Evarutil.new_evar env sigma ~principal:true ~store (EConstr.of_constr ty) end } end } @@ -384,7 +385,7 @@ let rename_hyp repl = let nctx = Environ.val_of_named_context nhyps in let instance = List.map (NamedDecl.get_id %> mkVar) hyps in Refine.refine ~unsafe:true { run = begin fun sigma -> - Evarutil.new_evar_instance nctx sigma nconcl ~principal:true ~store instance + Evarutil.new_evar_instance nctx sigma (EConstr.of_constr nconcl) ~principal:true ~store instance end } end } @@ -494,7 +495,7 @@ let rec mk_holes : type r s. _ -> r Sigma.t -> (s, r) Sigma.le -> _ -> (_, s) Si fun env sigma p -> function | [] -> Sigma ([], sigma, p) | arg :: rem -> - let Sigma (arg, sigma, q) = Evarutil.new_evar env sigma arg in + let Sigma (arg, sigma, q) = Evarutil.new_evar env sigma (EConstr.of_constr arg) in let Sigma (rem, sigma, r) = mk_holes env sigma (p +> q) rem in Sigma (arg :: rem, sigma, r) @@ -784,35 +785,37 @@ let make_change_arg c pats = { run = fun sigma -> Sigma.here (replace_vars (Id.Map.bindings pats) c) sigma } let check_types env sigma mayneedglobalcheck deep newc origc = - let t1 = Retyping.get_type_of env sigma (EConstr.of_constr newc) in + let t1 = Retyping.get_type_of env sigma newc in + let t1 = EConstr.of_constr t1 in if deep then begin - let t2 = Retyping.get_type_of env sigma (EConstr.of_constr origc) in + let t2 = Retyping.get_type_of env sigma origc in let sigma, t2 = Evarsolve.refresh_universes ~onlyalg:true (Some false) env sigma (EConstr.of_constr t2) in + let t2 = EConstr.of_constr t2 in let sigma, b = infer_conv ~pb:Reduction.CUMUL env sigma t1 t2 in if not b then if - isSort (whd_all env sigma (EConstr.of_constr t1)) && - isSort (whd_all env sigma (EConstr.of_constr t2)) + isSort (whd_all env sigma t1) && + isSort (whd_all env sigma t2) then (mayneedglobalcheck := true; sigma) else user_err ~hdr:"convert-check-hyp" (str "Types are incompatible.") else sigma end else - if not (isSort (whd_all env sigma (EConstr.of_constr t1))) then + if not (isSort (whd_all env sigma t1)) then user_err ~hdr:"convert-check-hyp" (str "Not a type.") else sigma (* Now we introduce different instances of the previous tacticals *) let change_and_check cv_pb mayneedglobalcheck deep t = { e_redfun = begin fun env sigma c -> - let c = EConstr.Unsafe.to_constr c in let Sigma (t', sigma, p) = t.run sigma in let sigma = Sigma.to_evar_map sigma in + let t' = EConstr.of_constr t' in let sigma = check_types env sigma mayneedglobalcheck deep t' c in let sigma, b = infer_conv ~pb:cv_pb env sigma t' c in if not b then user_err ~hdr:"convert-check-hyp" (str "Not convertible."); - Sigma.Unsafe.of_pair (t', sigma) + Sigma.Unsafe.of_pair (EConstr.Unsafe.to_constr t', sigma) end } (* Use cumulativity only if changing the conclusion not a subterm *) @@ -1240,8 +1243,8 @@ let cut c = (** Backward compat: normalize [c]. *) let c = if normalize_cut then local_strong whd_betaiota sigma (EConstr.of_constr c) else c in Refine.refine ~unsafe:true { run = begin fun h -> - let Sigma (f, h, p) = Evarutil.new_evar ~principal:true env h (mkArrow c (Vars.lift 1 concl)) in - let Sigma (x, h, q) = Evarutil.new_evar env h c in + let Sigma (f, h, p) = Evarutil.new_evar ~principal:true env h (EConstr.of_constr (mkArrow c (Vars.lift 1 concl))) in + let Sigma (x, h, q) = Evarutil.new_evar env h (EConstr.of_constr c) in let f = mkLetIn (Name id, x, c, mkApp (Vars.lift 1 f, [|mkRel 1|])) in Sigma (f, h, p +> q) end } @@ -1913,8 +1916,8 @@ let cut_and_apply c = let env = Tacmach.New.pf_env gl in Refine.refine { run = begin fun sigma -> let typ = mkProd (Anonymous, c2, concl) in - let Sigma (f, sigma, p) = Evarutil.new_evar env sigma typ in - let Sigma (x, sigma, q) = Evarutil.new_evar env sigma c1 in + let Sigma (f, sigma, p) = Evarutil.new_evar env sigma (EConstr.of_constr typ) in + let Sigma (x, sigma, q) = Evarutil.new_evar env sigma (EConstr.of_constr c1) in let ans = mkApp (f, [|mkApp (c, [|x|])|]) in Sigma (ans, sigma, p +> q) end } @@ -1983,7 +1986,7 @@ let assumption = if only_eq then (sigma, Constr.equal t concl) else let env = Proofview.Goal.env gl in - infer_conv env sigma t concl + infer_conv env sigma (EConstr.of_constr t) (EConstr.of_constr concl) in if is_same_type then (Proofview.Unsafe.tclEVARS sigma) <*> @@ -2078,7 +2081,7 @@ let clear_body ids = in check <*> Refine.refine ~unsafe:true { run = begin fun sigma -> - Evarutil.new_evar env sigma ~principal:true concl + Evarutil.new_evar env sigma ~principal:true (EConstr.of_constr concl) end } end } @@ -2131,7 +2134,7 @@ let apply_type newcl args = Refine.refine { run = begin fun sigma -> let newcl = nf_betaiota (Sigma.to_evar_map sigma) (EConstr.of_constr newcl) (* As in former Logic.refine *) in let Sigma (ev, sigma, p) = - Evarutil.new_evar env sigma ~principal:true ~store newcl in + Evarutil.new_evar env sigma ~principal:true ~store (EConstr.of_constr newcl) in Sigma (applist (ev, args), sigma, p) end } end } @@ -2151,7 +2154,7 @@ let bring_hyps hyps = let args = Array.of_list (Context.Named.to_instance hyps) in Refine.refine { run = begin fun sigma -> let Sigma (ev, sigma, p) = - Evarutil.new_evar env sigma ~principal:true ~store newcl in + Evarutil.new_evar env sigma ~principal:true ~store (EConstr.of_constr newcl) in Sigma (mkApp (ev, args), sigma, p) end } end } @@ -2677,11 +2680,11 @@ let mkletin_goal env sigma store with_eq dep (id,lastlhyp,ccl,c) ty = let eq = applist (eq,args) in let refl = applist (refl, [t;mkVar id]) in let newenv = insert_before [LocalAssum (heq,eq); decl] lastlhyp env in - let Sigma (x, sigma, r) = new_evar newenv sigma ~principal:true ~store ccl in + let Sigma (x, sigma, r) = new_evar newenv sigma ~principal:true ~store (EConstr.of_constr ccl) in Sigma (mkNamedLetIn id c t (mkNamedLetIn heq refl eq x), sigma, p +> q +> r) | None -> let newenv = insert_before [decl] lastlhyp env in - let Sigma (x, sigma, p) = new_evar newenv sigma ~principal:true ~store ccl in + let Sigma (x, sigma, p) = new_evar newenv sigma ~principal:true ~store (EConstr.of_constr ccl) in Sigma (mkNamedLetIn id c t x, sigma, p) let letin_tac with_eq id c ty occs = @@ -2862,7 +2865,7 @@ let new_generalize_gen_let lconstr = in let tac = Refine.refine { run = begin fun sigma -> - let Sigma (ev, sigma, p) = Evarutil.new_evar env sigma ~principal:true newcl in + let Sigma (ev, sigma, p) = Evarutil.new_evar env sigma ~principal:true (EConstr.of_constr newcl) in Sigma ((applist (ev, args)), sigma, p) end } in @@ -3549,7 +3552,7 @@ let make_abstract_generalize env id typ concl dep ctx body c eqs args refls = (* Abstract by the extension of the context *) let genctyp = it_mkProd_or_LetIn genarg ctx in (* The goal will become this product. *) - let Sigma (genc, sigma, p) = Evarutil.new_evar env sigma ~principal:true genctyp in + let Sigma (genc, sigma, p) = Evarutil.new_evar env sigma ~principal:true (EConstr.of_constr genctyp) in (* Apply the old arguments giving the proper instantiation of the hyp *) let instc = mkApp (genc, Array.of_list args) in (* Then apply to the original instantiated hyp. *) @@ -3755,7 +3758,7 @@ let specialize_eqs id gl = | _ -> if in_eqs then acc, in_eqs, ctx, ty else - let e = e_new_evar (push_rel_context ctx env) evars t in + let e = e_new_evar (push_rel_context ctx env) evars (EConstr.of_constr t) in aux false (LocalDef (na,e,t) :: ctx) (mkApp (lift 1 acc, [| mkRel 1 |])) b) | t -> acc, in_eqs, ctx, ty in diff --git a/toplevel/classes.ml b/toplevel/classes.ml index 26e29540c..1055d28b6 100644 --- a/toplevel/classes.ml +++ b/toplevel/classes.ml @@ -193,7 +193,7 @@ let new_instance ?(abstract=false) ?(global=false) ?(refine= !refine_instance) p let t = it_mkProd_or_LetIn ty_constr (ctx' @ ctx) in nf t in - Pretyping.check_evars env Evd.empty !evars termtype; + Pretyping.check_evars env Evd.empty !evars (EConstr.of_constr termtype); let pl, ctx = Evd.universe_context ?names:pl !evars in let cst = Declare.declare_constant ~internal:Declare.InternalTacticRequest id (ParameterEntry @@ -289,7 +289,7 @@ let new_instance ?(abstract=false) ?(global=false) ?(refine= !refine_instance) p let evm, nf = Evarutil.nf_evar_map_universes !evars in let termtype = nf termtype in let _ = (* Check that the type is free of evars now. *) - Pretyping.check_evars env Evd.empty evm termtype + Pretyping.check_evars env Evd.empty evm (EConstr.of_constr termtype) in let term = Option.map nf term in if not (Evd.has_undefined evm) && not (Option.is_empty term) then @@ -363,7 +363,7 @@ let context poly l = let _, ((env', fullctx), impls) = interp_context_evars env evars l in let subst = Evarutil.evd_comb0 Evarutil.nf_evars_and_universes evars in let fullctx = Context.Rel.map subst fullctx in - let ce t = Pretyping.check_evars env Evd.empty !evars t in + let ce t = Pretyping.check_evars env Evd.empty !evars (EConstr.of_constr t) in let () = List.iter (fun decl -> Context.Rel.Declaration.iter_constr ce decl) fullctx in let ctx = try named_of_rel_context fullctx diff --git a/toplevel/command.ml b/toplevel/command.ml index c6dd3eb99..dbf256ba8 100644 --- a/toplevel/command.ml +++ b/toplevel/command.ml @@ -618,10 +618,10 @@ let interp_mutual_inductive (paramsl,indl) notations poly prv finite = let ctx_params = Context.Rel.map nf ctx_params in let evd = !evdref in let pl, uctx = Evd.universe_context ?names:pl evd in - List.iter (check_evars env_params Evd.empty evd) arities; - Context.Rel.iter (check_evars env0 Evd.empty evd) ctx_params; + List.iter (fun c -> check_evars env_params Evd.empty evd (EConstr.of_constr c)) arities; + Context.Rel.iter (fun c -> check_evars env0 Evd.empty evd (EConstr.of_constr c)) ctx_params; List.iter (fun (_,ctyps,_) -> - List.iter (check_evars env_ar_params Evd.empty evd) ctyps) + List.iter (fun c -> check_evars env_ar_params Evd.empty evd (EConstr.of_constr c)) ctyps) constructors; (* Build the inductive entries *) @@ -1031,7 +1031,7 @@ let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation = mkApp (Universes.constr_of_global (delayed_force fix_sub_ref), [| argtyp ; wf_rel ; Evarutil.e_new_evar env evdref - ~src:(Loc.ghost, Evar_kinds.QuestionMark (Evar_kinds.Define false)) wf_proof; + ~src:(Loc.ghost, Evar_kinds.QuestionMark (Evar_kinds.Define false)) (EConstr.of_constr wf_proof); prop |]) in let def = Typing.e_solve_evars env evdref (EConstr.of_constr def) in diff --git a/toplevel/record.ml b/toplevel/record.ml index dc49c5385..929505716 100644 --- a/toplevel/record.ml +++ b/toplevel/record.ml @@ -163,7 +163,7 @@ let typecheck_params_and_fields def id pl t ps nots fs = let evars, nf = Evarutil.nf_evars_and_universes evars in let newps = Context.Rel.map nf newps in let newfs = Context.Rel.map nf newfs in - let ce t = Pretyping.check_evars env0 Evd.empty evars t in + let ce t = Pretyping.check_evars env0 Evd.empty evars (EConstr.of_constr t) in List.iter (iter_constr ce) (List.rev newps); List.iter (iter_constr ce) (List.rev newfs); Evd.universe_context ?names:pl evars, nf arity, template, imps, newps, impls, newfs @@ -364,13 +364,13 @@ let structure_signature ctx = | [decl] -> let env = Environ.empty_named_context_val in let evm = Sigma.Unsafe.of_evar_map evm in - let Sigma (_, evm, _) = Evarutil.new_pure_evar env evm (RelDecl.get_type decl) in + let Sigma (_, evm, _) = Evarutil.new_pure_evar env evm (EConstr.of_constr (RelDecl.get_type decl)) in let evm = Sigma.to_evar_map evm in evm | decl::tl -> let env = Environ.empty_named_context_val in let evm = Sigma.Unsafe.of_evar_map evm in - let Sigma (ev, evm, _) = Evarutil.new_pure_evar env evm (RelDecl.get_type decl) in + let Sigma (ev, evm, _) = Evarutil.new_pure_evar env evm (EConstr.of_constr (RelDecl.get_type decl)) in let evm = Sigma.to_evar_map evm in let new_tl = Util.List.map_i (fun pos decl -> diff --git a/toplevel/vernacentries.ml b/toplevel/vernacentries.ml index 33b96bc37..01b15407b 100644 --- a/toplevel/vernacentries.ml +++ b/toplevel/vernacentries.ml @@ -462,7 +462,7 @@ let start_proof_and_print k l hook = let c, _, ctx = Pfedit.build_by_tactic env (Evd.evar_universe_context sigma) concl (Tacticals.New.tclCOMPLETE tac) - in Evd.set_universe_context sigma ctx, c + in Evd.set_universe_context sigma ctx, EConstr.of_constr c with Logic_monad.TacticFailure e when Logic.catchable_exception e -> error "The statement obligations could not be resolved \ automatically, write a statement definition first." -- cgit v1.2.3 From c2855a3387be134d1220f301574b743572a94239 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Thu, 10 Nov 2016 11:39:27 +0100 Subject: Unification API using EConstr. --- engine/termops.ml | 27 +- engine/termops.mli | 6 +- ltac/rewrite.ml | 11 +- ltac/tacinterp.ml | 2 +- plugins/decl_mode/decl_proof_instr.ml | 2 +- plugins/extraction/extraction.ml | 3 + plugins/ssrmatching/ssrmatching.ml4 | 2 +- pretyping/evarconv.ml | 10 +- pretyping/evarsolve.ml | 8 +- pretyping/evarsolve.mli | 2 +- pretyping/find_subterm.ml | 2 +- pretyping/find_subterm.mli | 2 +- pretyping/inductiveops.ml | 2 +- pretyping/pretype_errors.ml | 16 +- pretyping/pretype_errors.mli | 30 +- pretyping/reductionops.ml | 2 +- pretyping/reductionops.mli | 2 +- pretyping/unification.ml | 504 +++++++++++++++++++--------------- pretyping/unification.mli | 32 ++- proofs/clenv.ml | 8 +- proofs/clenvtac.ml | 2 +- proofs/evar_refiner.ml | 2 +- proofs/goal.ml | 2 +- proofs/refine.ml | 2 +- tactics/equality.ml | 2 +- tactics/inv.ml | 4 +- tactics/tactics.ml | 29 +- toplevel/himsg.ml | 23 +- 28 files changed, 427 insertions(+), 312 deletions(-) diff --git a/engine/termops.ml b/engine/termops.ml index 16e2c04c8..c1198e05a 100644 --- a/engine/termops.ml +++ b/engine/termops.ml @@ -164,9 +164,19 @@ let rel_list n m = in reln [] 1 +let local_assum (na, t) = + let open RelDecl in + let inj = EConstr.Unsafe.to_constr in + LocalAssum (na, inj t) + +let local_def (na, b, t) = + let open RelDecl in + let inj = EConstr.Unsafe.to_constr in + LocalDef (na, inj b, inj t) + let push_rel_assum (x,t) env = let open RelDecl in - push_rel (LocalAssum (x,t)) env + push_rel (local_assum (x,t)) env let push_rels_assum assums = let open RelDecl in @@ -278,14 +288,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) + (mkApp (f1,extras), restl1, f2, l2) (* [map_constr_with_named_binders g f l c] maps [f l] on the immediate subterms of [c]; it carries an extra data [l] (typically a name @@ -321,16 +332,6 @@ let map_constr_with_named_binders g f l c = match kind_of_term c with time being almost those of the ML representation (except for (co-)fixpoint) *) -let local_assum (na, t) = - let open RelDecl in - let inj = EConstr.Unsafe.to_constr in - LocalAssum (na, inj t) - -let local_def (na, b, t) = - let open RelDecl in - let inj = EConstr.Unsafe.to_constr in - LocalDef (na, inj b, inj t) - let fold_rec_types g (lna,typarray,_) e = let open EConstr in let ctxt = Array.map2_i (fun i na t -> local_assum (na, Vars.lift i t)) lna typarray in diff --git a/engine/termops.mli b/engine/termops.mli index 2b66c88a7..df3fdd124 100644 --- a/engine/termops.mli +++ b/engine/termops.mli @@ -29,7 +29,7 @@ val print_rel_context : env -> std_ppcmds val print_env : env -> std_ppcmds (** about contexts *) -val push_rel_assum : Name.t * types -> env -> env +val push_rel_assum : Name.t * EConstr.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 @@ -208,8 +208,8 @@ val decompose_app_vect : Evd.evar_map -> EConstr.t -> constr * constr array val adjust_app_list_size : constr -> constr list -> constr -> constr list -> (constr * constr list * constr * constr list) -val adjust_app_array_size : constr -> constr array -> constr -> constr array -> - (constr * constr array * constr * constr array) +val adjust_app_array_size : EConstr.constr -> EConstr.constr array -> EConstr.constr -> EConstr.constr array -> + (EConstr.constr * EConstr.constr array * EConstr.constr * EConstr.constr array) (** name contexts *) type names_context = Name.t list diff --git a/ltac/rewrite.ml b/ltac/rewrite.ml index ccd45756a..4d65b4c69 100644 --- a/ltac/rewrite.ml +++ b/ltac/rewrite.ml @@ -711,7 +711,7 @@ let symmetry env sort rew = let unify_eqn (car, rel, prf, c1, c2, holes, sort) l2r flags env (sigma, cstrs) by t = try let left = if l2r then c1 else c2 in - let sigma = Unification.w_unify ~flags env sigma CONV left t in + let sigma = Unification.w_unify ~flags env sigma CONV (EConstr.of_constr left) (EConstr.of_constr t) in let sigma = Typeclasses.resolve_typeclasses ~filter:(no_constraints cstrs) ~fail:true env sigma in let evd = solve_remaining_by env sigma holes by in @@ -738,7 +738,7 @@ let unify_abs (car, rel, prf, c1, c2) l2r sort env (sigma, cstrs) t = basically an eq_constr, except when preexisting evars occur in either the lemma or the goal, in which case the eq_constr also solved this evars *) - let sigma = Unification.w_unify ~flags:rewrite_unif_flags env sigma CONV left t in + let sigma = Unification.w_unify ~flags:rewrite_unif_flags env sigma CONV (EConstr.of_constr left) (EConstr.of_constr t) in let rew_evars = sigma, cstrs in let rew_prf = RewPrf (rel, prf) in let rew = { rew_car = car; rew_from = c1; rew_to = c2; rew_prf; rew_evars; } in @@ -1423,7 +1423,7 @@ module Strategies = error "fold: the term is not unfoldable !" in try - let sigma = Unification.w_unify env sigma CONV ~flags:(Unification.elim_flags ()) unfolded t in + let sigma = Unification.w_unify env sigma CONV ~flags:(Unification.elim_flags ()) (EConstr.of_constr unfolded) (EConstr.of_constr t) in let c' = Evarutil.nf_evar sigma c in state, Success { rew_car = ty; rew_from = t; rew_to = c'; rew_prf = RewCast DEFAULTcast; @@ -2045,15 +2045,16 @@ let unification_rewrite l2r c1 c2 sigma prf car rel but env = in the context *) Unification.w_unify_to_subterm ~flags:rewrite_unif_flags - env sigma ((if l2r then c1 else c2),but) + env sigma (EConstr.of_constr (if l2r then c1 else c2),EConstr.of_constr but) with | ex when Pretype_errors.precatchable_exception ex -> (* ~flags:(true,true) to make Ring work (since it really exploits conversion) *) Unification.w_unify_to_subterm ~flags:rewrite_conv_unif_flags - env sigma ((if l2r then c1 else c2),but) + env sigma (EConstr.of_constr (if l2r then c1 else c2),EConstr.of_constr but) in + let c' = EConstr.Unsafe.to_constr c' in let nf c = Evarutil.nf_evar sigma c in let c1 = if l2r then nf c' else nf c1 and c2 = if l2r then nf c2 else nf c' diff --git a/ltac/tacinterp.ml b/ltac/tacinterp.ml index 85d4ac06e..d8c933912 100644 --- a/ltac/tacinterp.ml +++ b/ltac/tacinterp.ml @@ -1780,7 +1780,7 @@ and interp_atomic ist tac : unit Proofview.tactic = (TacLetTac(na,c,clp,b,eqpat)) (Tacticals.New.tclWITHHOLES false (*in hope of a future "eset/epose"*) (let_pat_tac b (interp_name ist env sigma na) - ((sigma,sigma'),c) clp eqpat) sigma') + ((sigma,sigma'),EConstr.of_constr c) clp eqpat) sigma') end } (* Derived basic tactics *) diff --git a/plugins/decl_mode/decl_proof_instr.ml b/plugins/decl_mode/decl_proof_instr.ml index 5e16d2da0..44cd22c8b 100644 --- a/plugins/decl_mode/decl_proof_instr.ml +++ b/plugins/decl_mode/decl_proof_instr.ml @@ -408,7 +408,7 @@ let find_subsubgoal c ctyp skip submetas gls = try let unifier = Unification.w_unify env se.se_evd Reduction.CUMUL - ~flags:(Unification.elim_flags ()) ctyp se.se_type in + ~flags:(Unification.elim_flags ()) (EConstr.of_constr ctyp) (EConstr.of_constr se.se_type) in if n <= 0 then {se with se_evd=meta_assign se.se_meta diff --git a/plugins/extraction/extraction.ml b/plugins/extraction/extraction.ml index 42a8cac69..0c4fa7055 100644 --- a/plugins/extraction/extraction.ml +++ b/plugins/extraction/extraction.ml @@ -99,6 +99,9 @@ let is_info_scheme env t = match flag_of_type env t with | (Info, TypeScheme) -> true | _ -> false +let push_rel_assum (n, t) env = + Environ.push_rel (LocalAssum (n, t)) env + (*s [type_sign] gernerates a signature aimed at treating a type application. *) let rec type_sign env c = diff --git a/plugins/ssrmatching/ssrmatching.ml4 b/plugins/ssrmatching/ssrmatching.ml4 index 18aeca6fa..308fb414e 100644 --- a/plugins/ssrmatching/ssrmatching.ml4 +++ b/plugins/ssrmatching/ssrmatching.ml4 @@ -332,7 +332,7 @@ let flags_FO = (Unification.default_no_delta_unify_flags ()).Unification.resolve_evars } let unif_FO env ise p c = - Unification.w_unify env ise Reduction.CONV ~flags:flags_FO p c + Unification.w_unify env ise Reduction.CONV ~flags:flags_FO (EConstr.of_constr p) (EConstr.of_constr c) (* Perform evar substitution in main term and prune substitution. *) let nf_open_term sigma0 ise c = diff --git a/pretyping/evarconv.ml b/pretyping/evarconv.ml index 683b33b89..3b420347b 100644 --- a/pretyping/evarconv.ml +++ b/pretyping/evarconv.ml @@ -390,7 +390,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty let t2 = EConstr.of_constr (nf_evar evd (EConstr.Unsafe.to_constr tM)) (** FIXME *) in let t2 = solve_pattern_eqn env evd l1' t2 in solve_simple_eqn (evar_conv_x ts) env evd - (position_problem on_left pbty,ev,EConstr.of_constr t2) + (position_problem on_left pbty,ev,t2) in let consume_stack on_left (termF,skF) (termO,skO) evd = let switch f a b = if on_left then f a b else f b a in @@ -1216,7 +1216,7 @@ let error_cannot_unify env evd pb ?reason t1 t2 = let check_problems_are_solved env evd = match snd (extract_all_conv_pbs evd) with - | (pbty,env,t1,t2) as pb::_ -> error_cannot_unify env evd pb t1 t2 + | (pbty,env,t1,t2) as pb::_ -> error_cannot_unify env evd pb (EConstr.of_constr t1) (EConstr.of_constr t2) | _ -> () let max_undefined_with_candidates evd = @@ -1276,7 +1276,9 @@ let consider_remaining_unif_problems env let rec aux evd pbs progress stuck = match pbs with | (pbty,env,t1,t2 as pb) :: pbs -> - (match apply_conversion_problem_heuristic ts env evd pbty (EConstr.of_constr t1) (EConstr.of_constr t2) with + let t1 = EConstr.of_constr t1 in + let t2 = EConstr.of_constr t2 in + (match apply_conversion_problem_heuristic ts env evd pbty t1 t2 with | Success evd' -> let (evd', rest) = extract_all_conv_pbs evd' in begin match rest with @@ -1292,6 +1294,8 @@ let consider_remaining_unif_problems env match stuck with | [] -> (* We're finished *) evd | (pbty,env,t1,t2 as pb) :: _ -> + let t1 = EConstr.of_constr t1 in + let t2 = EConstr.of_constr t2 in (* There remains stuck problems *) error_cannot_unify env evd pb t1 t2 in diff --git a/pretyping/evarsolve.ml b/pretyping/evarsolve.ml index 3bcea4cee..b1fc7cbe9 100644 --- a/pretyping/evarsolve.ml +++ b/pretyping/evarsolve.ml @@ -147,17 +147,17 @@ let recheck_applications conv_algo env evdref t = | App (f, args) -> let () = aux env f in let fty = Retyping.get_type_of env !evdref f in - let argsty = Array.map (fun x -> aux env x; Retyping.get_type_of env !evdref x) args in + let argsty = Array.map (fun x -> aux env x; EConstr.of_constr (Retyping.get_type_of env !evdref x)) args in let rec aux i ty = if i < Array.length argsty then match EConstr.kind !evdref (EConstr.of_constr (whd_all env !evdref ty)) with | Prod (na, dom, codom) -> - (match conv_algo env !evdref Reduction.CUMUL (EConstr.of_constr argsty.(i)) dom with + (match conv_algo env !evdref Reduction.CUMUL argsty.(i) dom with | Success evd -> evdref := evd; aux (succ i) (Vars.subst1 args.(i) codom) | UnifFailure (evd, reason) -> - Pretype_errors.error_cannot_unify env evd ~reason (argsty.(i), EConstr.Unsafe.to_constr dom)) - | _ -> raise (IllTypedInstance (env, ty, EConstr.of_constr argsty.(i))) + Pretype_errors.error_cannot_unify env evd ~reason (argsty.(i), dom)) + | _ -> raise (IllTypedInstance (env, ty, argsty.(i))) else () in aux 0 (EConstr.of_constr fty) | _ -> diff --git a/pretyping/evarsolve.mli b/pretyping/evarsolve.mli index 23cb245e0..b83147514 100644 --- a/pretyping/evarsolve.mli +++ b/pretyping/evarsolve.mli @@ -63,7 +63,7 @@ val is_unification_pattern_evar : env -> evar_map -> existential -> constr list val is_unification_pattern : env * int -> evar_map -> constr -> constr list -> constr -> constr list option -val solve_pattern_eqn : env -> evar_map -> constr list -> constr -> Constr.t +val solve_pattern_eqn : env -> evar_map -> constr list -> constr -> constr val noccur_evar : env -> evar_map -> Evar.t -> constr -> bool diff --git a/pretyping/find_subterm.ml b/pretyping/find_subterm.ml index 2b243d5b9..15409f2b8 100644 --- a/pretyping/find_subterm.ml +++ b/pretyping/find_subterm.ml @@ -85,7 +85,7 @@ let map_named_declaration_with_hyploc f hyploc acc decl = exception SubtermUnificationError of subterm_unification_error -exception NotUnifiable of (Constr.t * Constr.t * unification_error) option +exception NotUnifiable of (EConstr.t * EConstr.t * unification_error) option type 'a testing_function = { match_fun : 'a -> EConstr.constr -> 'a; diff --git a/pretyping/find_subterm.mli b/pretyping/find_subterm.mli index e7f0da93f..c7db84e3c 100644 --- a/pretyping/find_subterm.mli +++ b/pretyping/find_subterm.mli @@ -15,7 +15,7 @@ open Environ (** Finding subterms, possibly up to some unification function, possibly at some given occurrences *) -exception NotUnifiable of (constr * constr * unification_error) option +exception NotUnifiable of (EConstr.constr * EConstr.constr * unification_error) option exception SubtermUnificationError of subterm_unification_error diff --git a/pretyping/inductiveops.ml b/pretyping/inductiveops.ml index a93f2846b..e30ba21fd 100644 --- a/pretyping/inductiveops.ml +++ b/pretyping/inductiveops.ml @@ -504,7 +504,7 @@ let is_predicate_explicitly_dep env sigma pred arsign = let pv' = EConstr.of_constr (whd_all env sigma pval) in match EConstr.kind sigma pv', arsign with | Lambda (na,t,b), (LocalAssum _)::arsign -> - srec (push_rel_assum (na, EConstr.Unsafe.to_constr t) env) b arsign + srec (push_rel_assum (na, t) env) b arsign | Lambda (na,_,t), _ -> (* The following code has an impact on the introduction names diff --git a/pretyping/pretype_errors.ml b/pretyping/pretype_errors.ml index c14d81505..14b25ab36 100644 --- a/pretyping/pretype_errors.ml +++ b/pretyping/pretype_errors.ml @@ -29,7 +29,7 @@ type position = (Id.t * Locus.hyp_location_flag) option type position_reporting = (position * int) * EConstr.t -type subterm_unification_error = bool * position_reporting * position_reporting * (constr * constr * unification_error) option +type subterm_unification_error = bool * position_reporting * position_reporting * (EConstr.constr * EConstr.constr * unification_error) option type pretype_error = (* Old Case *) @@ -37,17 +37,17 @@ type pretype_error = (* Type inference unification *) | ActualTypeNotCoercible of unsafe_judgment * types * unification_error (* Tactic unification *) - | UnifOccurCheck of existential_key * constr + | UnifOccurCheck of existential_key * EConstr.constr | UnsolvableImplicit of existential_key * Evd.unsolvability_explanation option - | CannotUnify of constr * constr * unification_error option - | CannotUnifyLocal of constr * constr * constr + | CannotUnify of EConstr.constr * EConstr.constr * unification_error option + | CannotUnifyLocal of EConstr.constr * EConstr.constr * EConstr.constr | CannotUnifyBindingType of constr * constr | CannotGeneralize of constr - | NoOccurrenceFound of constr * Id.t option - | CannotFindWellTypedAbstraction of constr * EConstr.constr list * (env * type_error) option - | WrongAbstractionType of Name.t * constr * types * types + | NoOccurrenceFound of EConstr.constr * Id.t option + | CannotFindWellTypedAbstraction of EConstr.constr * EConstr.constr list * (env * type_error) option + | WrongAbstractionType of Name.t * EConstr.constr * EConstr.types * EConstr.types | AbstractionOverMeta of Name.t * Name.t - | NonLinearUnification of Name.t * constr + | NonLinearUnification of Name.t * EConstr.constr (* Pretyping *) | VarNotFound of Id.t | UnexpectedType of constr * constr diff --git a/pretyping/pretype_errors.mli b/pretyping/pretype_errors.mli index 217deda4d..2e707a0ff 100644 --- a/pretyping/pretype_errors.mli +++ b/pretyping/pretype_errors.mli @@ -30,7 +30,7 @@ type position = (Id.t * Locus.hyp_location_flag) option type position_reporting = (position * int) * EConstr.t -type subterm_unification_error = bool * position_reporting * position_reporting * (constr * constr * unification_error) option +type subterm_unification_error = bool * position_reporting * position_reporting * (EConstr.constr * EConstr.constr * unification_error) option type pretype_error = (** Old Case *) @@ -38,17 +38,17 @@ type pretype_error = (** Type inference unification *) | ActualTypeNotCoercible of unsafe_judgment * types * unification_error (** Tactic Unification *) - | UnifOccurCheck of existential_key * constr + | UnifOccurCheck of existential_key * EConstr.constr | UnsolvableImplicit of existential_key * Evd.unsolvability_explanation option - | CannotUnify of constr * constr * unification_error option - | CannotUnifyLocal of constr * constr * constr + | CannotUnify of EConstr.constr * EConstr.constr * unification_error option + | CannotUnifyLocal of EConstr.constr * EConstr.constr * EConstr.constr | CannotUnifyBindingType of constr * constr | CannotGeneralize of constr - | NoOccurrenceFound of constr * Id.t option - | CannotFindWellTypedAbstraction of constr * EConstr.constr list * (env * type_error) option - | WrongAbstractionType of Name.t * constr * types * types + | NoOccurrenceFound of EConstr.constr * Id.t option + | CannotFindWellTypedAbstraction of EConstr.constr * EConstr.constr list * (env * type_error) option + | WrongAbstractionType of Name.t * EConstr.constr * EConstr.types * EConstr.types | AbstractionOverMeta of Name.t * Name.t - | NonLinearUnification of Name.t * constr + | NonLinearUnification of Name.t * EConstr.constr (** Pretyping *) | VarNotFound of Id.t | UnexpectedType of constr * constr @@ -94,32 +94,32 @@ val error_ill_typed_rec_body : val error_not_a_type : ?loc:Loc.t -> env -> Evd.evar_map -> unsafe_judgment -> 'b -val error_cannot_coerce : env -> Evd.evar_map -> constr * constr -> 'b +val error_cannot_coerce : env -> Evd.evar_map -> EConstr.constr * EConstr.constr -> 'b (** {6 Implicit arguments synthesis errors } *) -val error_occur_check : env -> Evd.evar_map -> existential_key -> constr -> 'b +val error_occur_check : env -> Evd.evar_map -> existential_key -> EConstr.constr -> 'b val error_unsolvable_implicit : ?loc:Loc.t -> env -> Evd.evar_map -> existential_key -> Evd.unsolvability_explanation option -> 'b val error_cannot_unify : ?loc:Loc.t -> env -> Evd.evar_map -> - ?reason:unification_error -> constr * constr -> 'b + ?reason:unification_error -> EConstr.constr * EConstr.constr -> 'b -val error_cannot_unify_local : env -> Evd.evar_map -> constr * constr * constr -> 'b +val error_cannot_unify_local : env -> Evd.evar_map -> EConstr.constr * EConstr.constr * EConstr.constr -> 'b val error_cannot_find_well_typed_abstraction : env -> Evd.evar_map -> - constr -> EConstr.constr list -> (env * type_error) option -> 'b + EConstr.constr -> EConstr.constr list -> (env * type_error) option -> 'b val error_wrong_abstraction_type : env -> Evd.evar_map -> - Name.t -> constr -> types -> types -> 'b + Name.t -> EConstr.constr -> EConstr.types -> EConstr.types -> 'b val error_abstraction_over_meta : env -> Evd.evar_map -> metavariable -> metavariable -> 'b val error_non_linear_unification : env -> Evd.evar_map -> - metavariable -> constr -> 'b + metavariable -> EConstr.constr -> 'b (** {6 Ml Case errors } *) diff --git a/pretyping/reductionops.ml b/pretyping/reductionops.ml index 0b97cd253..510417879 100644 --- a/pretyping/reductionops.ml +++ b/pretyping/reductionops.ml @@ -1203,7 +1203,7 @@ let whd_allnolet env = (* 4. Ad-hoc eta reduction, does not subsitute evars *) -let shrink_eta c = EConstr.Unsafe.to_constr (Stack.zip Evd.empty (local_whd_state_gen eta Evd.empty (c,Stack.empty))) +let shrink_eta c = Stack.zip Evd.empty (local_whd_state_gen eta Evd.empty (c,Stack.empty)) (* 5. Zeta Reduction Functions *) diff --git a/pretyping/reductionops.mli b/pretyping/reductionops.mli index 5e6a40786..c3b82729d 100644 --- a/pretyping/reductionops.mli +++ b/pretyping/reductionops.mli @@ -198,7 +198,7 @@ val whd_zeta_stack : local_stack_reduction_function val whd_zeta_state : local_state_reduction_function val whd_zeta : local_reduction_function -val shrink_eta : EConstr.t -> constr +val shrink_eta : EConstr.constr -> EConstr.constr (** Various reduction functions *) diff --git a/pretyping/unification.ml b/pretyping/unification.ml index 70aa0be6b..c5c19b49b 100644 --- a/pretyping/unification.ml +++ b/pretyping/unification.ml @@ -6,13 +6,16 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +module CVars = Vars + open CErrors open Pp open Util open Names open Term -open Vars open Termops +open EConstr +open Vars open Namegen open Environ open Evd @@ -30,6 +33,13 @@ open Locusops open Find_subterm open Sigma.Notations +type metabinding = (metavariable * EConstr.constr * (instance_constraint * instance_typing_status)) + +type subst0 = + (evar_map * + metabinding list * + (Environ.env * EConstr.existential * EConstr.t) list) + module RelDecl = Context.Rel.Declaration module NamedDecl = Context.Named.Declaration @@ -55,41 +65,44 @@ let _ = Goptions.declare_bool_option { } let occur_meta_or_undefined_evar evd c = - let rec occrec c = match kind_of_term c with + let rec occrec c = match EConstr.kind evd c with | Meta _ -> raise Occur - | Evar (ev,args) -> - (match evar_body (Evd.find evd ev) with - | Evar_defined c -> - occrec c; Array.iter occrec args - | Evar_empty -> raise Occur) - | _ -> Constr.iter occrec c + | Evar _ -> raise Occur + | _ -> EConstr.iter evd occrec c in try occrec c; false with Occur | Not_found -> true let occur_meta_evd sigma mv c = let rec occrec c = (* Note: evars are not instantiated by terms with metas *) - let c = whd_evar sigma (whd_meta sigma (EConstr.of_constr c)) in - match kind_of_term c with + let c = whd_meta sigma c in + let c = EConstr.of_constr c in + match EConstr.kind sigma c with | Meta mv' when Int.equal mv mv' -> raise Occur - | _ -> Constr.iter occrec c + | _ -> EConstr.iter sigma occrec c in try occrec c; false with Occur -> true (* if lname_typ is [xn,An;..;x1,A1] and l is a list of terms, gives [x1:A1]..[xn:An]c' such that c converts to ([x1:A1]..[xn:An]c' l) *) let abstract_scheme env evd c l lname_typ = + let open EConstr in + let mkLambda_name env (n,a,b) = + mkLambda (named_hd env (EConstr.Unsafe.to_constr a) n, a, b) + in List.fold_left2 (fun (t,evd) (locc,a) decl -> let na = RelDecl.get_name decl in let ta = RelDecl.get_type decl in - let na = match kind_of_term a with Var id -> Name id | _ -> na in + let ta = EConstr.of_constr ta in + let na = match EConstr.kind evd a with Var id -> Name id | _ -> na in (* [occur_meta ta] test removed for support of eelim/ecase but consequences are unclear... if occur_meta ta then error "cannot find a type for the generalisation" else *) - if occur_meta evd (EConstr.of_constr a) then mkLambda_name env (na,ta,t), evd + if occur_meta evd a then mkLambda_name env (na,ta,t), evd else - let t', evd' = Find_subterm.subst_closed_term_occ env evd locc (EConstr.of_constr a) (EConstr.of_constr t) in + let t', evd' = Find_subterm.subst_closed_term_occ env evd locc a t in + let t' = EConstr.of_constr t' in mkLambda_name env (na,ta,t'), evd') (c,evd) (List.rev l) @@ -98,16 +111,17 @@ let abstract_scheme env evd c l lname_typ = (* Precondition: resulting abstraction is expected to be of type [typ] *) let abstract_list_all env evd typ c l = - let ctxt,_ = splay_prod_n env evd (List.length l) (EConstr.of_constr typ) in + let ctxt,_ = splay_prod_n env evd (List.length l) typ in let l_with_all_occs = List.map (function a -> (LikeFirst,a)) l in let p,evd = abstract_scheme env evd c l_with_all_occs ctxt in let evd,typp = - try Typing.type_of env evd (EConstr.of_constr p) + try Typing.type_of env evd p with | UserError _ -> - error_cannot_find_well_typed_abstraction env evd p (List.map EConstr.of_constr l) None + error_cannot_find_well_typed_abstraction env evd p l None | Type_errors.TypeError (env',x) -> - error_cannot_find_well_typed_abstraction env evd p (List.map EConstr.of_constr l) (Some (env',x)) in + error_cannot_find_well_typed_abstraction env evd p l (Some (env',x)) in + let typp = EConstr.of_constr typp in evd,(p,typp) let set_occurrences_of_last_arg args = @@ -123,12 +137,12 @@ let abstract_list_all_with_dependencies env evd typ c l = let argoccs = set_occurrences_of_last_arg (Array.sub (snd ev') 0 n) in let evd,b = Evarconv.second_order_matching empty_transparent_state - env evd ev' argoccs (EConstr.of_constr c) in + env evd ev' argoccs c in if b then let p = nf_evar evd ev in evd, p else error_cannot_find_well_typed_abstraction env evd - (nf_evar evd c) l None + c l None (**) @@ -148,51 +162,53 @@ let extract_instance_status = function | CUMUL -> add_type_status (IsSubType, IsSuperType) | CONV -> add_type_status (Conv, Conv) -let rec subst_meta_instances bl c = - match kind_of_term c with +let rec subst_meta_instances sigma bl c = + match EConstr.kind sigma c with | Meta i -> let select (j,_,_) = Int.equal i j in (try pi2 (List.find select bl) with Not_found -> c) - | _ -> Constr.map (subst_meta_instances bl) c + | _ -> EConstr.map sigma (subst_meta_instances sigma bl) c (** [env] should be the context in which the metas live *) let pose_all_metas_as_evars env evd t = let evdref = ref evd in - let rec aux t = match kind_of_term t with + let rec aux t = match EConstr.kind !evdref t with | Meta mv -> (match Evd.meta_opt_fvalue !evdref mv with - | Some ({rebus=c},_) -> c + | Some ({rebus=c},_) -> EConstr.of_constr c | None -> let {rebus=ty;freemetas=mvs} = Evd.meta_ftype evd mv in + let ty = EConstr.of_constr ty in let ty = if Evd.Metaset.is_empty mvs then ty else aux ty in let src = Evd.evar_source_of_meta mv !evdref in - let ty = EConstr.of_constr ty in let ev = Evarutil.e_new_evar env evdref ~src ty in evdref := meta_assign mv (ev,(Conv,TypeNotProcessed)) !evdref; - ev) + EConstr.of_constr ev) | _ -> - Constr.map aux t in + EConstr.map !evdref aux t in let c = aux t in (* side-effect *) (!evdref, c) -let solve_pattern_eqn_array (env,nb) f l c (sigma,metasubst,evarsubst) = - match kind_of_term f with +let solve_pattern_eqn_array (env,nb) f l c (sigma,metasubst,evarsubst : subst0) = + let open EConstr in + let open Vars in + match EConstr.kind sigma f with | Meta k -> (* We enforce that the Meta does not depend on the [nb] extra assumptions added by unification to the context *) let env' = pop_rel_context nb env in let sigma,c = pose_all_metas_as_evars env' sigma c in - let c = solve_pattern_eqn env sigma (List.map EConstr.of_constr l) (EConstr.of_constr c) in + let c = solve_pattern_eqn env sigma l c in let pb = (Conv,TypeNotProcessed) in - if noccur_between 1 nb c then + if noccur_between sigma 1 nb c then sigma,(k,lift (-nb) c,pb)::metasubst,evarsubst else error_cannot_unify_local env sigma (applist (f, l),c,c) | Evar ev -> let env' = pop_rel_context nb env in let sigma,c = pose_all_metas_as_evars env' sigma c in - sigma,metasubst,(env,ev,solve_pattern_eqn env sigma (List.map EConstr.of_constr l) (EConstr.of_constr c))::evarsubst + sigma,metasubst,(env,ev,solve_pattern_eqn env sigma l c)::evarsubst | _ -> assert false let push d (env,n) = (push_rel_assum d env,n+1) @@ -456,15 +472,16 @@ let use_evars_pattern_unification flags = !global_pattern_unification_flag && flags.use_pattern_unification && Flags.version_strictly_greater Flags.V8_2 -let use_metas_pattern_unification flags nb l = +let use_metas_pattern_unification sigma flags nb l = + let open EConstr in !global_pattern_unification_flag && flags.use_pattern_unification || (Flags.version_less_or_equal Flags.V8_3 || flags.use_meta_bound_pattern_unification) && - Array.for_all (fun c -> isRel c && destRel c <= nb) l + Array.for_all (fun c -> isRel sigma c && destRel sigma c <= nb) l type key = | IsKey of CClosure.table_key - | IsProj of projection * constr + | IsProj of projection * EConstr.constr let expand_table_key env = function | ConstKey cst -> constant_opt_value_in env cst @@ -480,11 +497,11 @@ let unfold_projection env p stk = | None -> assert false) let expand_key ts env sigma = function - | Some (IsKey k) -> expand_table_key env k + | Some (IsKey k) -> Option.map EConstr.of_constr (expand_table_key env k) | Some (IsProj (p, c)) -> - let red = EConstr.Unsafe.to_constr (Stack.zip sigma (fst (whd_betaiota_deltazeta_for_iota_state ts env sigma - Cst_stack.empty (EConstr.of_constr c, unfold_projection env p [])))) - in if Term.eq_constr (mkProj (p, c)) red then None else Some red + let red = Stack.zip sigma (fst (whd_betaiota_deltazeta_for_iota_state ts env sigma + Cst_stack.empty (c, unfold_projection env p []))) + in if EConstr.eq_constr sigma (EConstr.mkProj (p, c)) red then None else Some red | None -> None @@ -497,9 +514,9 @@ type unirec_flags = { let subterm_restriction opt flags = not opt.at_top && flags.restrict_conv_on_strict_subterms -let key_of env b flags f = +let key_of env sigma b flags f = if subterm_restriction b flags then None else - match kind_of_term f with + match EConstr.kind sigma f with | Const (cst, u) when is_transparent env (ConstKey cst) && (Cpred.mem cst (snd flags.modulo_delta) || Environ.is_projection cst env) -> @@ -544,8 +561,8 @@ let oracle_order env cf1 cf2 = Some (Conv_oracle.oracle_order (fun x -> x) (Environ.oracle env) false (translate_key k1) (translate_key k2)) -let is_rigid_head flags t = - match kind_of_term t with +let is_rigid_head sigma flags t = + match EConstr.kind sigma t with | Const (cst,u) -> not (Cpred.mem cst (snd flags.modulo_delta)) | Ind (i,u) -> true | Construct _ -> true @@ -561,15 +578,15 @@ let force_eqs c = let constr_cmp pb sigma flags t u = let cstrs = - if pb == Reduction.CONV then Universes.eq_constr_universes t u - else Universes.leq_constr_universes t u + if pb == Reduction.CONV then EConstr.eq_constr_universes sigma t u + else EConstr.leq_constr_universes sigma t u in match cstrs with | Some cstrs -> begin try Evd.add_universe_constraints sigma cstrs, true with Univ.UniverseInconsistency _ -> sigma, false | Evd.UniversesDiffer -> - if is_rigid_head flags t then + if is_rigid_head sigma flags t then try Evd.add_universe_constraints sigma (force_eqs cstrs), true with Univ.UniverseInconsistency _ -> sigma, false else sigma, false @@ -578,46 +595,47 @@ let constr_cmp pb sigma flags t u = sigma, false let do_reduce ts (env, nb) sigma c = - EConstr.Unsafe.to_constr (Stack.zip sigma (fst (whd_betaiota_deltazeta_for_iota_state - ts env sigma Cst_stack.empty (EConstr.of_constr c, Stack.empty)))) + Stack.zip sigma (fst (whd_betaiota_deltazeta_for_iota_state + ts env sigma Cst_stack.empty (c, Stack.empty))) let use_full_betaiota flags = flags.modulo_betaiota && Flags.version_strictly_greater Flags.V8_3 -let isAllowedEvar flags c = match kind_of_term c with +let isAllowedEvar sigma flags c = match EConstr.kind sigma c with | Evar (evk,_) -> not (Evar.Set.mem evk flags.frozen_evars) | _ -> false -let subst_defined_metas_evars (bl,el) c = - let rec substrec c = match kind_of_term c with +let subst_defined_metas_evars sigma (bl,el) c = + let rec substrec c = match EConstr.kind sigma c with | Meta i -> let select (j,_,_) = Int.equal i j in substrec (pi2 (List.find select bl)) | Evar (evk,args) -> - let select (_,(evk',args'),_) = Evar.equal evk evk' && Array.equal Constr.equal args args' in + let select (_,(evk',args'),_) = Evar.equal evk evk' && Array.equal (EConstr.eq_constr sigma) args args' in (try substrec (pi3 (List.find select el)) - with Not_found -> Constr.map substrec c) - | _ -> Constr.map substrec c + with Not_found -> EConstr.map sigma substrec c) + | _ -> EConstr.map sigma substrec c in try Some (substrec c) with Not_found -> None -let check_compatibility env pbty flags (sigma,metasubst,evarsubst) tyM tyN = - match subst_defined_metas_evars (metasubst,[]) tyM with +let check_compatibility env pbty flags (sigma,metasubst,evarsubst : subst0) tyM tyN = + match subst_defined_metas_evars sigma (metasubst,[]) (EConstr.of_constr tyM) with | None -> sigma | Some m -> - match subst_defined_metas_evars (metasubst,[]) tyN with + match subst_defined_metas_evars sigma (metasubst,[]) (EConstr.of_constr tyN) with | None -> sigma | Some n -> - if is_ground_term sigma (EConstr.of_constr m) && is_ground_term sigma (EConstr.of_constr n) then - let sigma, b = infer_conv ~pb:pbty ~ts:flags.modulo_delta_types env sigma (EConstr.of_constr m) (EConstr.of_constr n) in + if is_ground_term sigma m && is_ground_term sigma n then + let sigma, b = infer_conv ~pb:pbty ~ts:flags.modulo_delta_types env sigma m n in if b then sigma else error_cannot_unify env sigma (m,n) else sigma -let rec is_neutral env ts t = - let (f, l) = decompose_appvect t in - match kind_of_term f with +let rec is_neutral env sigma ts t = + let open EConstr in + let (f, l) = decompose_app_vect sigma t in + match EConstr.kind sigma (EConstr.of_constr f) with | Const (c, u) -> not (Environ.evaluable_constant c env) || not (is_transparent env (ConstKey c)) || @@ -628,24 +646,25 @@ let rec is_neutral env ts t = not (Id.Pred.mem id (fst ts)) | Rel n -> true | Evar _ | Meta _ -> true - | Case (_, p, c, cl) -> is_neutral env ts c - | Proj (p, c) -> is_neutral env ts c + | Case (_, p, c, cl) -> is_neutral env sigma ts c + | Proj (p, c) -> is_neutral env sigma ts c | _ -> false -let is_eta_constructor_app env ts f l1 term = - match kind_of_term f with +let is_eta_constructor_app env sigma ts f l1 term = + match EConstr.kind sigma f with | Construct (((_, i as ind), j), u) when i == 0 && j == 1 -> let mib = lookup_mind (fst ind) env in (match mib.Declarations.mind_record with | Some (Some (_,exp,projs)) when mib.Declarations.mind_finite == Decl_kinds.BiFinite && Array.length projs == Array.length l1 - mib.Declarations.mind_nparams -> (** Check that the other term is neutral *) - is_neutral env ts term + is_neutral env sigma ts term | _ -> false) | _ -> false -let eta_constructor_app env f l1 term = - match kind_of_term f with +let eta_constructor_app env sigma f l1 term = + let open EConstr in + match EConstr.kind sigma f with | Construct (((_, i as ind), j), u) -> let mib = lookup_mind (fst ind) env in (match mib.Declarations.mind_record with @@ -658,15 +677,20 @@ let eta_constructor_app env f l1 term = | _ -> assert false) | _ -> assert false -let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb flags m n = - let rec unirec_rec (curenv,nb as curenvnb) pb opt ((sigma,metasubst,evarsubst) as substn) curm curn = - let cM = EConstr.Unsafe.to_constr (Evarutil.whd_head_evar sigma (EConstr.of_constr curm)) - and cN = EConstr.Unsafe.to_constr (Evarutil.whd_head_evar sigma (EConstr.of_constr curn)) in +let print_constr_env env c = + print_constr_env env (EConstr.Unsafe.to_constr c) + +let rec unify_0_with_initial_metas (sigma,ms,es as subst : subst0) conv_at_top env cv_pb flags m n = + let open EConstr in + let open Vars in + let rec unirec_rec (curenv,nb as curenvnb) pb opt ((sigma,metasubst,evarsubst) as substn : subst0) curm curn = + let cM = Evarutil.whd_head_evar sigma curm + and cN = Evarutil.whd_head_evar sigma curn in let () = if !debug_unification then - Feedback.msg_debug (Termops.print_constr_env curenv cM ++ str" ~= " ++ Termops.print_constr_env curenv cN) + Feedback.msg_debug (print_constr_env curenv cM ++ str" ~= " ++ print_constr_env curenv cN) in - match (kind_of_term cM,kind_of_term cN) with + match (EConstr.kind sigma cM, EConstr.kind sigma cN) with | Meta k1, Meta k2 -> if Int.equal k1 k2 then substn else let stM,stN = extract_instance_status pb in @@ -681,12 +705,12 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb if k2 < k1 then sigma,(k1,cN,stN)::metasubst,evarsubst else sigma,(k2,cM,stM)::metasubst,evarsubst | Meta k, _ - when not (dependent sigma (EConstr.of_constr cM) (EConstr.of_constr cN)) (* helps early trying alternatives *) -> + when not (dependent sigma cM cN) (* helps early trying alternatives *) -> let sigma = if opt.with_types && flags.check_applied_meta_types then (try let tyM = Typing.meta_type sigma k in - let tyN = get_type_of curenv ~lax:true sigma (EConstr.of_constr cN) in + let tyN = get_type_of curenv ~lax:true sigma cN in check_compatibility curenv CUMUL flags substn tyN tyM with RetypeError _ -> (* Renounce, maybe metas/evars prevents typing *) sigma) @@ -695,17 +719,17 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb (* Here we check that [cN] does not contain any local variables *) if Int.equal nb 0 then sigma,(k,cN,snd (extract_instance_status pb))::metasubst,evarsubst - else if noccur_between 1 nb cN then + else if noccur_between sigma 1 nb cN then (sigma, (k,lift (-nb) cN,snd (extract_instance_status pb))::metasubst, evarsubst) else error_cannot_unify_local curenv sigma (m,n,cN) | _, Meta k - when not (dependent sigma (EConstr.of_constr cN) (EConstr.of_constr cM)) (* helps early trying alternatives *) -> + when not (dependent sigma cN cM) (* helps early trying alternatives *) -> let sigma = if opt.with_types && flags.check_applied_meta_types then (try - let tyM = get_type_of curenv ~lax:true sigma (EConstr.of_constr cM) in + let tyM = get_type_of curenv ~lax:true sigma cM in let tyN = Typing.meta_type sigma k in check_compatibility curenv CUMUL flags substn tyM tyN with RetypeError _ -> @@ -715,7 +739,7 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb (* Here we check that [cM] does not contain any local variables *) if Int.equal nb 0 then (sigma,(k,cM,fst (extract_instance_status pb))::metasubst,evarsubst) - else if noccur_between 1 nb cM + else if noccur_between sigma 1 nb cM then (sigma,(k,lift (-nb) cM,fst (extract_instance_status pb))::metasubst, evarsubst) @@ -730,15 +754,15 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb sigma,metasubst,((curenv,ev,cN)::evarsubst) | Evar (evk,_ as ev), _ when not (Evar.Set.mem evk flags.frozen_evars) - && not (occur_evar sigma evk (EConstr.of_constr cN)) -> - let cmvars = free_rels sigma (EConstr.of_constr cM) and cnvars = free_rels sigma (EConstr.of_constr cN) in + && not (occur_evar sigma evk cN) -> + let cmvars = free_rels sigma cM and cnvars = free_rels sigma cN in if Int.Set.subset cnvars cmvars then sigma,metasubst,((curenv,ev,cN)::evarsubst) else error_cannot_unify_local curenv sigma (m,n,cN) | _, Evar (evk,_ as ev) when not (Evar.Set.mem evk flags.frozen_evars) - && not (occur_evar sigma evk (EConstr.of_constr cM)) -> - let cmvars = free_rels sigma (EConstr.of_constr cM) and cnvars = free_rels sigma (EConstr.of_constr cN) in + && not (occur_evar sigma evk cM) -> + let cmvars = free_rels sigma cM and cnvars = free_rels sigma cN in if Int.Set.subset cmvars cnvars then sigma,metasubst,((curenv,ev,cM)::evarsubst) else error_cannot_unify_local curenv sigma (m,n,cN) @@ -781,30 +805,30 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb | App (f1, l1), _ when flags.modulo_eta && (* This ensures cN is an evar, meta or irreducible constant/variable and not a constructor. *) - is_eta_constructor_app curenv flags.modulo_delta f1 l1 cN -> + is_eta_constructor_app curenv sigma flags.modulo_delta f1 l1 cN -> (try - let l1', l2' = eta_constructor_app curenv f1 l1 cN in + let l1', l2' = eta_constructor_app curenv sigma f1 l1 cN in let opt' = {opt with at_top = true; with_cs = false} in Array.fold_left2 (unirec_rec curenvnb CONV opt') substn l1' l2' with ex when precatchable_exception ex -> - match kind_of_term cN with + match EConstr.kind sigma cN with | App(f2,l2) when - (isMeta f2 && use_metas_pattern_unification flags nb l2 - || use_evars_pattern_unification flags && isAllowedEvar flags f2) -> + (isMeta sigma f2 && use_metas_pattern_unification sigma flags nb l2 + || use_evars_pattern_unification flags && isAllowedEvar sigma flags f2) -> unify_app_pattern false curenvnb pb opt substn cM f1 l1 cN f2 l2 | _ -> raise ex) | _, App (f2, l2) when flags.modulo_eta && - is_eta_constructor_app curenv flags.modulo_delta f2 l2 cM -> + is_eta_constructor_app curenv sigma flags.modulo_delta f2 l2 cM -> (try - let l2', l1' = eta_constructor_app curenv f2 l2 cM in + let l2', l1' = eta_constructor_app curenv sigma f2 l2 cM in let opt' = {opt with at_top = true; with_cs = false} in Array.fold_left2 (unirec_rec curenvnb CONV opt') substn l1' l2' with ex when precatchable_exception ex -> - match kind_of_term cM with + match EConstr.kind sigma cM with | App(f1,l1) when - (isMeta f1 && use_metas_pattern_unification flags nb l1 - || use_evars_pattern_unification flags && isAllowedEvar flags f1) -> + (isMeta sigma f1 && use_metas_pattern_unification sigma flags nb l1 + || use_evars_pattern_unification flags && isAllowedEvar sigma flags f1) -> unify_app_pattern true curenvnb pb opt substn cM f1 l1 cN f2 l2 | _ -> raise ex) @@ -819,13 +843,13 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb reduce curenvnb pb opt substn cM cN) | App (f1,l1), _ when - (isMeta f1 && use_metas_pattern_unification flags nb l1 - || use_evars_pattern_unification flags && isAllowedEvar flags f1) -> + (isMeta sigma f1 && use_metas_pattern_unification sigma flags nb l1 + || use_evars_pattern_unification flags && isAllowedEvar sigma flags f1) -> unify_app_pattern true curenvnb pb opt substn cM f1 l1 cN cN [||] | _, App (f2,l2) when - (isMeta f2 && use_metas_pattern_unification flags nb l2 - || use_evars_pattern_unification flags && isAllowedEvar flags f2) -> + (isMeta sigma f2 && use_metas_pattern_unification sigma flags nb l2 + || use_evars_pattern_unification flags && isAllowedEvar sigma flags f2) -> unify_app_pattern false curenvnb pb opt substn cM cM [||] cN f2 l2 | App (f1,l1), App (f2,l2) -> @@ -840,32 +864,32 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb | _ -> unify_not_same_head curenvnb pb opt substn cM cN - and unify_app_pattern dir curenvnb pb opt substn cM f1 l1 cN f2 l2 = + and unify_app_pattern dir curenvnb pb opt (sigma, _, _ as substn) cM f1 l1 cN f2 l2 = let f, l, t = if dir then f1, l1, cN else f2, l2, cM in - match is_unification_pattern curenvnb sigma (EConstr.of_constr f) (Array.map_to_list EConstr.of_constr l) (EConstr.of_constr t) with + match is_unification_pattern curenvnb sigma f (Array.to_list l) t with | None -> - (match kind_of_term t with + (match EConstr.kind sigma t with | App (f',l') -> if dir then unify_app curenvnb pb opt substn cM f1 l1 t f' l' else unify_app curenvnb pb opt substn t f' l' cN f2 l2 | Proj _ -> unify_app curenvnb pb opt substn cM f1 l1 cN f2 l2 | _ -> unify_not_same_head curenvnb pb opt substn cM cN) | Some l -> - solve_pattern_eqn_array curenvnb f (List.map EConstr.Unsafe.to_constr l) t substn + solve_pattern_eqn_array curenvnb f l t substn - and unify_app (curenv, nb as curenvnb) pb opt (sigma, metas, evars as substn) cM f1 l1 cN f2 l2 = + and unify_app (curenv, nb as curenvnb) pb opt (sigma, metas, evars as substn : subst0) cM f1 l1 cN f2 l2 = try let needs_expansion p c' = - match kind_of_term c' with + match EConstr.kind sigma c' with | Meta _ -> true | Evar _ -> true | Const (c, u) -> Constant.equal c (Projection.constant p) | _ -> false in let expand_proj c c' l = - match kind_of_term c with + match EConstr.kind sigma c with | Proj (p, t) when not (Projection.unfolded p) && needs_expansion p c' -> - (try destApp (Retyping.expand_projection curenv sigma p (EConstr.of_constr t) (Array.map_to_list EConstr.of_constr l)) + (try destApp sigma (EConstr.of_constr (Retyping.expand_projection curenv sigma p t (Array.to_list l))) with RetypeError _ -> (** Unification can be called on ill-typed terms, due to FO and eta in particular, fail gracefully in that case *) (c, l)) @@ -890,8 +914,10 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb and unify_same_proj (curenv, nb as curenvnb) cv_pb opt substn c1 c2 = let substn = unirec_rec curenvnb CONV opt substn c1 c2 in try (* Force unification of the types to fill in parameters *) - let ty1 = get_type_of curenv ~lax:true sigma (EConstr.of_constr c1) in - let ty2 = get_type_of curenv ~lax:true sigma (EConstr.of_constr c2) in + let ty1 = get_type_of curenv ~lax:true sigma c1 in + let ty2 = get_type_of curenv ~lax:true sigma c2 in + let ty1 = EConstr.of_constr ty1 in + let ty2 = EConstr.of_constr ty2 in unify_0_with_initial_metas substn true curenv cv_pb { flags with modulo_conv_on_closed_terms = Some full_transparent_state; modulo_delta = full_transparent_state; @@ -900,7 +926,7 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb ty1 ty2 with RetypeError _ -> substn - and unify_not_same_head curenvnb pb opt (sigma, metas, evars as substn) cM cN = + and unify_not_same_head curenvnb pb opt (sigma, metas, evars as substn : subst0) cM cN = try canonical_projections curenvnb pb opt cM cN substn with ex when precatchable_exception ex -> let sigma', b = constr_cmp cv_pb sigma flags cM cN in @@ -909,24 +935,24 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb try reduce curenvnb pb opt substn cM cN with ex when precatchable_exception ex -> let (f1,l1) = - match kind_of_term cM with App (f,l) -> (f,l) | _ -> (cM,[||]) in + match EConstr.kind sigma cM with App (f,l) -> (f,l) | _ -> (cM,[||]) in let (f2,l2) = - match kind_of_term cN with App (f,l) -> (f,l) | _ -> (cN,[||]) in + match EConstr.kind sigma cN with App (f,l) -> (f,l) | _ -> (cN,[||]) in expand curenvnb pb opt substn cM f1 l1 cN f2 l2 and reduce curenvnb pb opt (sigma, metas, evars as substn) cM cN = if use_full_betaiota flags && not (subterm_restriction opt flags) then let cM' = do_reduce flags.modulo_delta curenvnb sigma cM in - if not (Term.eq_constr cM cM') then + if not (EConstr.eq_constr sigma cM cM') then unirec_rec curenvnb pb opt substn cM' cN else let cN' = do_reduce flags.modulo_delta curenvnb sigma cN in - if not (Term.eq_constr cN cN') then + if not (EConstr.eq_constr sigma cN cN') then unirec_rec curenvnb pb opt substn cM cN' else error_cannot_unify (fst curenvnb) sigma (cM,cN) else error_cannot_unify (fst curenvnb) sigma (cM,cN) - and expand (curenv,_ as curenvnb) pb opt (sigma,metasubst,evarsubst as substn) cM f1 l1 cN f2 l2 = + and expand (curenv,_ as curenvnb) pb opt (sigma,metasubst,evarsubst as substn : subst0) cM f1 l1 cN f2 l2 = let res = (* Try full conversion on meta-free terms. *) (* Back to 1995 (later on called trivial_unify in 2002), the @@ -945,24 +971,23 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb | None -> None | Some convflags -> let subst = ((if flags.use_metas_eagerly_in_conv_on_closed_terms then metasubst else ms), (if flags.use_evars_eagerly_in_conv_on_closed_terms then evarsubst else es)) in - match subst_defined_metas_evars subst cM with + match subst_defined_metas_evars sigma subst cM with | None -> (* some undefined Metas in cM *) None | Some m1 -> - match subst_defined_metas_evars subst cN with + match subst_defined_metas_evars sigma subst cN with | None -> (* some undefined Metas in cN *) None | Some n1 -> (* No subterm restriction there, too much incompatibilities *) let sigma = if opt.with_types then try (* Ensure we call conversion on terms of the same type *) - let tyM = get_type_of curenv ~lax:true sigma (EConstr.of_constr m1) in - let tyN = get_type_of curenv ~lax:true sigma (EConstr.of_constr n1) in + let tyM = get_type_of curenv ~lax:true sigma m1 in + let tyN = get_type_of curenv ~lax:true sigma n1 in check_compatibility curenv CUMUL flags substn tyM tyN with RetypeError _ -> (* Renounce, maybe metas/evars prevents typing *) sigma else sigma in - let m1 = EConstr.of_constr m1 and n1 = EConstr.of_constr n1 in let sigma, b = infer_conv ~pb ~ts:convflags curenv sigma m1 n1 in if b then Some (sigma, metasubst, evarsubst) else @@ -973,40 +998,40 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb match res with | Some substn -> substn | None -> - let cf1 = key_of curenv opt flags f1 and cf2 = key_of curenv opt flags f2 in + let cf1 = key_of curenv sigma opt flags f1 and cf2 = key_of curenv sigma opt flags f2 in match oracle_order curenv cf1 cf2 with | None -> error_cannot_unify curenv sigma (cM,cN) | Some true -> (match expand_key flags.modulo_delta curenv sigma cf1 with | Some c -> unirec_rec curenvnb pb opt substn - (whd_betaiotazeta sigma (EConstr.of_constr (mkApp(c,l1)))) cN + (EConstr.of_constr (whd_betaiotazeta sigma (mkApp(c,l1)))) cN | None -> (match expand_key flags.modulo_delta curenv sigma cf2 with | Some c -> unirec_rec curenvnb pb opt substn cM - (whd_betaiotazeta sigma (EConstr.of_constr (mkApp(c,l2)))) + (EConstr.of_constr (whd_betaiotazeta sigma (mkApp(c,l2)))) | None -> error_cannot_unify curenv sigma (cM,cN))) | Some false -> (match expand_key flags.modulo_delta curenv sigma cf2 with | Some c -> unirec_rec curenvnb pb opt substn cM - (whd_betaiotazeta sigma (EConstr.of_constr (mkApp(c,l2)))) + (EConstr.of_constr (whd_betaiotazeta sigma (mkApp(c,l2)))) | None -> (match expand_key flags.modulo_delta curenv sigma cf1 with | Some c -> unirec_rec curenvnb pb opt substn - (whd_betaiotazeta sigma (EConstr.of_constr (mkApp(c,l1)))) cN + (EConstr.of_constr (whd_betaiotazeta sigma (mkApp(c,l1)))) cN | None -> error_cannot_unify curenv sigma (cM,cN))) and canonical_projections (curenv, _ as curenvnb) pb opt cM cN (sigma,_,_ as substn) = let f1 () = - if isApp cM then - let f1l1 = whd_nored_state sigma (EConstr.of_constr cM,Stack.empty) in + if isApp sigma cM then + let f1l1 = whd_nored_state sigma (cM,Stack.empty) in if is_open_canonical_projection curenv sigma f1l1 then - let f2l2 = whd_nored_state sigma (EConstr.of_constr cN,Stack.empty) in + let f2l2 = whd_nored_state sigma (cN,Stack.empty) in solve_canonical_projection curenvnb pb opt cM f1l1 cN f2l2 substn else error_cannot_unify (fst curenvnb) sigma (cM,cN) else error_cannot_unify (fst curenvnb) sigma (cM,cN) @@ -1019,10 +1044,10 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb error_cannot_unify (fst curenvnb) sigma (cM,cN) else try f1 () with e when precatchable_exception e -> - if isApp cN then - let f2l2 = whd_nored_state sigma (EConstr.of_constr cN, Stack.empty) in + if isApp sigma cN then + let f2l2 = whd_nored_state sigma (cN, Stack.empty) in if is_open_canonical_projection curenv sigma f2l2 then - let f1l1 = whd_nored_state sigma (EConstr.of_constr cM, Stack.empty) in + let f1l1 = whd_nored_state sigma (cM, Stack.empty) in solve_canonical_projection curenvnb pb opt cN f2l2 cM f1l1 substn else error_cannot_unify (fst curenvnb) sigma (cM,cN) else error_cannot_unify (fst curenvnb) sigma (cM,cN) @@ -1038,26 +1063,25 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb List.fold_left (fun (evd,ks,m) b -> if match n with Some n -> Int.equal m n | None -> false then - (evd,EConstr.Unsafe.to_constr t2::ks, m-1) + (evd,t2::ks, m-1) else let mv = new_meta () in - let evd' = meta_declare mv (substl ks b) evd in + let evd' = meta_declare mv (EConstr.Unsafe.to_constr (substl ks b)) evd in (evd', mkMeta mv :: ks, m - 1)) - (sigma,[],List.length bs) (List.map EConstr.Unsafe.to_constr bs) + (sigma,[],List.length bs) bs in try let opt' = {opt with with_types = false} in - let inj = EConstr.Unsafe.to_constr in let (substn,_,_) = Reductionops.Stack.fold2 - (fun s u1 u -> unirec_rec curenvnb pb opt' s (inj u1) (substl ks (inj u))) + (fun s u1 u -> unirec_rec curenvnb pb opt' s u1 (substl ks u)) (evd,ms,es) us2 us in let (substn,_,_) = Reductionops.Stack.fold2 - (fun s u1 u -> unirec_rec curenvnb pb opt' s (inj u1) (substl ks (inj u))) + (fun s u1 u -> unirec_rec curenvnb pb opt' s u1 (substl ks u)) substn params1 params in - let (substn,_,_) = Reductionops.Stack.fold2 (fun s u1 u2 -> unirec_rec curenvnb pb opt' s (inj u1) (inj u2)) substn ts ts1 in - let app = mkApp (EConstr.Unsafe.to_constr c, Array.rev_of_list ks) in + let (substn,_,_) = Reductionops.Stack.fold2 (fun s u1 u2 -> unirec_rec curenvnb pb opt' s u1 u2) substn ts ts1 in + let app = mkApp (c, Array.rev_of_list ks) in (* let substn = unirec_rec curenvnb pb b false substn t cN in *) - unirec_rec curenvnb pb opt' substn (EConstr.Unsafe.to_constr c1) app + unirec_rec curenvnb pb opt' substn c1 app with Invalid_argument "Reductionops.Stack.fold2" -> error_cannot_unify (fst curenvnb) sigma (cM,cN) else error_cannot_unify (fst curenvnb) sigma (cM,cN) @@ -1073,7 +1097,7 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb None else let sigma, b = match flags.modulo_conv_on_closed_terms with - | Some convflags -> infer_conv ~pb:cv_pb ~ts:convflags env sigma (EConstr.of_constr m) (EConstr.of_constr n) + | Some convflags -> infer_conv ~pb:cv_pb ~ts:convflags env sigma m n | _ -> constr_cmp cv_pb sigma flags m n in if b then Some sigma else if (match flags.modulo_conv_on_closed_terms, flags.modulo_delta with @@ -1101,7 +1125,9 @@ let right = false let rec unify_with_eta keptside flags env sigma c1 c2 = (* Question: try whd_all on ci if not two lambdas? *) - match kind_of_term c1, kind_of_term c2 with + let open EConstr in + let open Vars in + match EConstr.kind sigma c1, EConstr.kind sigma c2 with | (Lambda (na,t1,c1'), Lambda (_,t2,c2')) -> let env' = push_rel_assum (na,t1) env in let sigma,metas,evars = unify_0 env sigma CONV flags t1 t2 in @@ -1205,32 +1231,39 @@ let merge_instances env sigma flags st1 st2 c1 c2 = * since other metavars might also need to be resolved. *) let applyHead env (type r) (evd : r Sigma.t) n c = + let open EConstr in + let open Vars in let rec apprec : type s. _ -> _ -> _ -> (r, s) Sigma.le -> s Sigma.t -> (constr, r) Sigma.sigma = fun n c cty p evd -> if Int.equal n 0 then Sigma (c, evd, p) else - match kind_of_term (whd_all env (Sigma.to_evar_map evd) (EConstr.of_constr cty)) with + let sigma = Sigma.to_evar_map evd in + match EConstr.kind sigma (EConstr.of_constr (whd_all env sigma cty)) with | Prod (_,c1,c2) -> - let Sigma (evar, evd', q) = Evarutil.new_evar env evd ~src:(Loc.ghost,Evar_kinds.GoalEvar) (EConstr.of_constr c1) in + let Sigma (evar, evd', q) = Evarutil.new_evar env evd ~src:(Loc.ghost,Evar_kinds.GoalEvar) c1 in + let evar = EConstr.of_constr evar in apprec (n-1) (mkApp(c,[|evar|])) (subst1 evar c2) (p +> q) evd' | _ -> error "Apply_Head_Then" in - apprec n c (Typing.unsafe_type_of env (Sigma.to_evar_map evd) (EConstr.of_constr c)) Sigma.refl evd + apprec n c (EConstr.of_constr (Typing.unsafe_type_of env (Sigma.to_evar_map evd) c)) Sigma.refl evd -let is_mimick_head ts f = - match kind_of_term f with +let is_mimick_head sigma ts f = + match EConstr.kind sigma f with | Const (c,u) -> not (CClosure.is_transparent_constant ts c) | Var id -> not (CClosure.is_transparent_variable ts id) | (Rel _|Construct _|Ind _) -> true | _ -> false +let make_judge c t = + make_judge (EConstr.Unsafe.to_constr c) (EConstr.Unsafe.to_constr t) + let try_to_coerce env evd c cty tycon = let j = make_judge c cty in - let (evd',j') = inh_conv_coerce_rigid_to true Loc.ghost env evd j (EConstr.of_constr tycon) in + let (evd',j') = inh_conv_coerce_rigid_to true Loc.ghost env evd j tycon in let evd' = Evarconv.consider_remaining_unif_problems env evd' in let evd' = Evd.map_metas_fvalue (nf_evar evd') evd' in - (evd',j'.uj_val) + (evd',EConstr.of_constr j'.uj_val) let w_coerce_to_type env evd c cty mvty = let evd,tycon = pose_all_metas_as_evars env evd mvty in @@ -1239,19 +1272,19 @@ let w_coerce_to_type env evd c cty mvty = (* inh_conv_coerce_rigid_to should have reasoned modulo reduction but there are cases where it though it was not rigid (like in fst (nat,nat)) and stops while it could have seen that it is rigid *) - let cty = Tacred.hnf_constr env evd (EConstr.of_constr cty) in - try_to_coerce env evd c cty tycon + let cty = Tacred.hnf_constr env evd cty in + try_to_coerce env evd c (EConstr.of_constr cty) tycon let w_coerce env evd mv c = - let cty = get_type_of env evd (EConstr.of_constr c) in + let cty = get_type_of env evd c in let mvty = Typing.meta_type evd mv in - w_coerce_to_type env evd c cty mvty + w_coerce_to_type env evd c (EConstr.of_constr cty) (EConstr.of_constr mvty) let unify_to_type env sigma flags c status u = - let sigma, c = refresh_universes (Some false) env sigma (EConstr.of_constr c) in + let sigma, c = refresh_universes (Some false) env sigma c in let t = get_type_of env sigma (EConstr.of_constr (nf_meta sigma c)) in let t = nf_betaiota sigma (EConstr.of_constr (nf_meta sigma t)) in - unify_0 env sigma CUMUL flags t u + unify_0 env sigma CUMUL flags (EConstr.of_constr t) (EConstr.of_constr u) let unify_type env sigma flags mv status c = let mvty = Typing.meta_type sigma mv in @@ -1274,9 +1307,10 @@ let order_metas metas = (* Solve an equation ?n[x1=u1..xn=un] = t where ?n is an evar *) let solve_simple_evar_eqn ts env evd ev rhs = - match solve_simple_eqn (Evarconv.evar_conv_x ts) env evd (None,ev,EConstr.of_constr rhs) with + let open EConstr in + match solve_simple_eqn (Evarconv.evar_conv_x ts) env evd (None,ev,rhs) with | UnifFailure (evd,reason) -> - error_cannot_unify env evd ~reason (EConstr.Unsafe.to_constr (EConstr.mkEvar ev),rhs); + error_cannot_unify env evd ~reason (mkEvar ev,rhs); | Success evd -> Evarconv.consider_remaining_unif_problems env evd @@ -1284,25 +1318,27 @@ let solve_simple_evar_eqn ts env evd ev rhs = or in evars, possibly generating new unification problems; if [b] is true, unification of types of metas is required *) -let w_merge env with_types flags (evd,metas,evars) = +let w_merge env with_types flags (evd,metas,evars : subst0) = + let open EConstr in + let open Vars in let rec w_merge_rec evd metas evars eqns = (* Process evars *) match evars with | (curenv,(evk,_ as ev),rhs)::evars' -> if Evd.is_defined evd evk then - let v = Evd.existential_value evd ev in + let v = mkEvar ev in let (evd,metas',evars'') = unify_0 curenv evd CONV flags rhs v in w_merge_rec evd (metas'@metas) (evars''@evars') eqns else begin (* This can make rhs' ill-typed if metas are *) - let rhs' = subst_meta_instances metas rhs in - match kind_of_term rhs with - | App (f,cl) when occur_meta evd (EConstr.of_constr rhs') -> - if occur_evar evd evk (EConstr.of_constr rhs') then + let rhs' = subst_meta_instances evd metas rhs in + match EConstr.kind evd rhs with + | App (f,cl) when occur_meta evd rhs' -> + if occur_evar evd evk rhs' then error_occur_check curenv evd evk rhs'; - if is_mimick_head flags.modulo_delta f then + if is_mimick_head evd flags.modulo_delta f then let evd' = mimick_undefined_evar evd flags f (Array.length cl) evk in (* let evd' = Evarconv.consider_remaining_unif_problems env evd' in *) @@ -1310,14 +1346,14 @@ let w_merge env with_types flags (evd,metas,evars) = else let evd' = let evd', rhs'' = pose_all_metas_as_evars curenv evd rhs' in - try solve_simple_evar_eqn flags.modulo_delta_types curenv evd' (fst ev, Array.map EConstr.of_constr (snd ev)) rhs'' + try solve_simple_evar_eqn flags.modulo_delta_types curenv evd' ev rhs'' with Retyping.RetypeError _ -> error_cannot_unify curenv evd' (mkEvar ev,rhs'') in w_merge_rec evd' metas evars' eqns | _ -> let evd', rhs'' = pose_all_metas_as_evars curenv evd rhs' in let evd' = - try solve_simple_evar_eqn flags.modulo_delta_types curenv evd' (fst ev, Array.map EConstr.of_constr (snd ev)) rhs'' + try solve_simple_evar_eqn flags.modulo_delta_types curenv evd' ev rhs'' with Retyping.RetypeError _ -> error_cannot_unify curenv evd' (mkEvar ev, rhs'') in w_merge_rec evd' metas evars' eqns @@ -1343,20 +1379,20 @@ let w_merge env with_types flags (evd,metas,evars) = if meta_defined evd mv then let {rebus=c'},(status',_) = meta_fvalue evd mv in let (take_left,st,(evd,metas',evars')) = - merge_instances env evd flags status' status c' c + merge_instances env evd flags status' status (EConstr.of_constr c') c in let evd' = if take_left then evd - else meta_reassign mv (c,(st,TypeProcessed)) evd + else meta_reassign mv (EConstr.Unsafe.to_constr c,(st,TypeProcessed)) evd in w_merge_rec evd' (metas'@metas@metas'') (evars'@evars'') eqns else let evd' = if occur_meta_evd evd mv c then - if isMetaOf mv (whd_all env evd (EConstr.of_constr c)) then evd + if isMetaOf mv (whd_all env evd c) then evd else error_cannot_unify env evd (mkMeta mv,c) else - meta_assign mv (c,(status,TypeProcessed)) evd in + meta_assign mv (EConstr.Unsafe.to_constr c,(status,TypeProcessed)) evd in w_merge_rec evd' (metas''@metas) evars'' eqns | [] -> (* Process type eqns *) @@ -1382,17 +1418,17 @@ let w_merge env with_types flags (evd,metas,evars) = let evd' = Sigma.to_evar_map evd' in let (evd'',mc,ec) = unify_0 sp_env evd' CUMUL flags - (get_type_of sp_env evd' (EConstr.of_constr c)) ev.evar_concl in + (EConstr.of_constr (get_type_of sp_env evd' c)) (EConstr.of_constr ev.evar_concl) in let evd''' = w_merge_rec evd'' mc ec [] in if evd' == evd''' - then Evd.define sp c evd''' - else Evd.define sp (Evarutil.nf_evar evd''' c) evd''' in + then Evd.define sp (EConstr.Unsafe.to_constr c) evd''' + else Evd.define sp (Evarutil.nf_evar evd''' (EConstr.Unsafe.to_constr c)) evd''' in let check_types evd = let metas = Evd.meta_list evd in let eqns = List.fold_left (fun acc (mv, b) -> match b with - | Clval (n, (t, (c, TypeNotProcessed)), v) -> (mv, c, t.rebus) :: acc + | Clval (n, (t, (c, TypeNotProcessed)), v) -> (mv, c, EConstr.of_constr t.rebus) :: acc | _ -> acc) [] metas in w_merge_rec evd [] [] eqns in @@ -1404,6 +1440,11 @@ let w_merge env with_types flags (evd,metas,evars) = if with_types then check_types res else res +let retract_coercible_metas evd = + let (metas, evd) = retract_coercible_metas evd in + let map (mv, c, st) = (mv, EConstr.of_constr c, st) in + (List.map map metas, evd) + let w_unify_meta_types env ?(flags=default_unify_flags ()) evd = let metas,evd = retract_coercible_metas evd in w_merge env true flags.merge_unify_flags (evd,metas,[]) @@ -1419,19 +1460,23 @@ let w_unify_meta_types env ?(flags=default_unify_flags ()) evd = types of metavars are unifiable with the types of their instances *) let head_app sigma m = - EConstr.Unsafe.to_constr (fst (whd_nored_state sigma (EConstr.of_constr m, Stack.empty))) + fst (whd_nored_state sigma (m, Stack.empty)) + +let isEvar_or_Meta sigma c = match EConstr.kind sigma c with +| Evar _ | Meta _ -> true +| _ -> false let check_types env flags (sigma,_,_ as subst) m n = - if isEvar_or_Meta (head_app sigma m) then + if isEvar_or_Meta sigma (head_app sigma m) then unify_0_with_initial_metas subst true env CUMUL flags - (get_type_of env sigma (EConstr.of_constr n)) - (get_type_of env sigma (EConstr.of_constr m)) - else if isEvar_or_Meta (head_app sigma n) then + (EConstr.of_constr (get_type_of env sigma n)) + (EConstr.of_constr (get_type_of env sigma m)) + else if isEvar_or_Meta sigma (head_app sigma n) then unify_0_with_initial_metas subst true env CUMUL flags - (get_type_of env sigma (EConstr.of_constr m)) - (get_type_of env sigma (EConstr.of_constr n)) + (EConstr.of_constr (get_type_of env sigma m)) + (EConstr.of_constr (get_type_of env sigma n)) else subst let try_resolve_typeclasses env evd flag m n = @@ -1453,6 +1498,11 @@ let w_unify_core_0 env evd with_types cv_pb flags m n = let w_typed_unify env evd = w_unify_core_0 env evd true let w_typed_unify_array env evd flags f1 l1 f2 l2 = + let open EConstr in + let f1 = EConstr.of_constr f1 in + let f2 = EConstr.of_constr f2 in + let l1 = Array.map EConstr.of_constr l1 in + let l2 = Array.map EConstr.of_constr l2 in let f1,l1,f2,l2 = adjust_app_array_size f1 l1 f2 l2 in let (mc1,evd') = retract_coercible_metas evd in let fold_subst subst m n = unify_0_with_initial_metas subst true env CONV flags.core_unify_flags m n in @@ -1479,7 +1529,8 @@ let iter_fail f a = contexts, with evars, and possibly with occurrences *) let indirectly_dependent sigma c d decls = - not (isVar c) && + let open EConstr in + not (isVar sigma c) && (* This test is not needed if the original term is a variable, but it is needed otherwise, as e.g. when abstracting over "2" in "forall H:0=2, H=H:>(0=1+1) -> 0=2." where there is now obvious @@ -1493,7 +1544,8 @@ let finish_evar_resolution ?(flags=Pretyping.all_and_fail_flags) env current_sig let current_sigma = Sigma.to_evar_map current_sigma in let sigma = Pretyping.solve_remaining_evars flags env current_sigma pending in let sigma, subst = nf_univ_variables sigma in - Sigma.Unsafe.of_pair (subst_univs_constr subst (nf_evar sigma c), sigma) + let c = EConstr.Unsafe.to_constr c in + Sigma.Unsafe.of_pair (EConstr.of_constr (CVars.subst_univs_constr subst (nf_evar sigma c)), sigma) let default_matching_core_flags sigma = let ts = Names.full_transparent_state in { @@ -1538,6 +1590,7 @@ let default_matching_flags (sigma,_) = exception PatternNotFound let make_pattern_test from_prefix_of_ind is_correct_type env sigma (pending,c) = + let open EConstr in let flags = if from_prefix_of_ind then let flags = default_matching_flags pending in @@ -1545,7 +1598,7 @@ let make_pattern_test from_prefix_of_ind is_correct_type env sigma (pending,c) = modulo_conv_on_closed_terms = Some Names.full_transparent_state; restrict_conv_on_strict_subterms = true } } else default_matching_flags pending in - let n = List.length (snd (decompose_app c)) in + let n = Array.length (snd (decompose_app_vect sigma c)) in let matching_fun _ t = let open EConstr in try @@ -1560,8 +1613,9 @@ let make_pattern_test from_prefix_of_ind is_correct_type env sigma (pending,c) = else applist (t,l1), l2 else t, [] in - let sigma = w_typed_unify env sigma Reduction.CONV flags c (EConstr.Unsafe.to_constr t') in + let sigma = w_typed_unify env sigma Reduction.CONV flags c t' in let ty = Retyping.get_type_of env sigma t in + let ty = EConstr.of_constr ty in if not (is_correct_type ty) then raise (NotUnifiable None); Some(sigma, t, l2) with @@ -1582,19 +1636,20 @@ let make_pattern_test from_prefix_of_ind is_correct_type env sigma (pending,c) = (fun test -> match test.testing_state with | None -> None | Some (sigma,_,l) -> - let c = applist (nf_evar sigma (local_strong whd_meta sigma (EConstr.of_constr c)), List.map (EConstr.to_constr sigma) l) in + let c = applist (EConstr.of_constr (nf_evar sigma (local_strong whd_meta sigma c)), l) in let univs, subst = nf_univ_variables sigma in - Some (sigma,subst_univs_constr subst c)) + Some (sigma,EConstr.of_constr (CVars.subst_univs_constr subst (EConstr.Unsafe.to_constr c)))) let make_eq_test env evd c = let out cstr = - match cstr.last_found with None -> None | _ -> Some (cstr.testing_state, EConstr.Unsafe.to_constr c) + match cstr.last_found with None -> None | _ -> Some (cstr.testing_state, c) in (make_eq_univs_test env evd c, out) let make_abstraction_core name (test,out) env sigma c ty occs check_occs concl = let id = - let t = match ty with Some t -> t | None -> get_type_of env sigma (EConstr.of_constr c) in + let ty = Option.map EConstr.Unsafe.to_constr ty in + let t = match ty with Some t -> t | None -> get_type_of env sigma c in let x = id_of_name_using_hdchar (Global.env()) t name in let ids = ids_of_named_context (named_context env) in if name == Anonymous then next_ident_away_in_goal x ids else @@ -1634,7 +1689,7 @@ let make_abstraction_core name (test,out) env sigma c ty occs check_occs concl = | NoOccurrences -> concl | occ -> let occ = if likefirst then LikeFirst else AtOccs occ in - replace_term_occ_modulo sigma occ test mkvarid (EConstr.of_constr concl) + EConstr.of_constr (replace_term_occ_modulo sigma occ test mkvarid concl) in let lastlhyp = if List.is_empty depdecls then None else Some (NamedDecl.get_id (List.last depdecls)) in @@ -1660,6 +1715,8 @@ let make_abstraction_core name (test,out) env sigma c ty occs check_occs concl = type prefix_of_inductive_support_flag = bool +type pending_constr = Evd.pending * constr + type abstraction_request = | AbstractPattern of prefix_of_inductive_support_flag * (types -> bool) * Name.t * pending_constr * clause * bool | AbstractExact of Name.t * constr * types option * clause * bool @@ -1678,7 +1735,7 @@ let make_abstraction env evd ccl abs = env evd (snd c) None occs check_occs ccl | AbstractExact (name,c,ty,occs,check_occs) -> make_abstraction_core name - (make_eq_test env evd (EConstr.of_constr c)) + (make_eq_test env evd c) env evd c ty occs check_occs ccl let keyed_unify env evd kop = @@ -1688,7 +1745,7 @@ let keyed_unify env evd kop = | None -> fun _ -> true | Some kop -> fun cl -> - let kc = Keys.constr_key cl in + let kc = Keys.constr_key (EConstr.to_constr evd cl) in match kc with | None -> false | Some kc -> Keys.equiv_keys kop kc @@ -1697,23 +1754,25 @@ let keyed_unify env evd kop = Unifies [cl] to every subterm of [op] until it finds a match. Fails if no match is found *) let w_unify_to_subterm env evd ?(flags=default_unify_flags ()) (op,cl) = + let open EConstr in + let open Vars in let bestexn = ref None in - let kop = Keys.constr_key op in + let kop = Keys.constr_key (EConstr.to_constr evd op) in let rec matchrec cl = - let cl = strip_outer_cast evd (EConstr.of_constr cl) in + let cl = EConstr.of_constr (strip_outer_cast evd cl) in (try - if closed0 cl && not (isEvar cl) && keyed_unify env evd kop cl then + if closed0 evd cl && not (isEvar evd cl) && keyed_unify env evd kop cl then (try if !keyed_unification then - let f1, l1 = decompose_app_vect evd (EConstr.of_constr op) in - let f2, l2 = decompose_app_vect evd (EConstr.of_constr cl) in + let f1, l1 = decompose_app_vect evd op in + let f2, l2 = decompose_app_vect evd cl in w_typed_unify_array env evd flags f1 l1 f2 l2,cl else w_typed_unify env evd CONV flags op cl,cl with ex when Pretype_errors.unsatisfiable_exception ex -> bestexn := Some ex; error "Unsat") else error "Bound 1" with ex when precatchable_exception ex -> - (match kind_of_term cl with + (match EConstr.kind evd cl with | App (f,args) -> let n = Array.length args in assert (n>0); @@ -1772,9 +1831,11 @@ let w_unify_to_subterm env evd ?(flags=default_unify_flags ()) (op,cl) = Unifies [cl] to every subterm of [op] and return all the matches. Fails if no match is found *) let w_unify_to_subterm_all env evd ?(flags=default_unify_flags ()) (op,cl) = + let open EConstr in + let open Vars in let return a b = let (evd,c as a) = a () in - if List.exists (fun (evd',c') -> Term.eq_constr c c') b then b else a :: b + if List.exists (fun (evd',c') -> EConstr.eq_constr evd' c c') b then b else a :: b in let fail str _ = error str in let bind f g a = @@ -1793,12 +1854,13 @@ let w_unify_to_subterm_all env evd ?(flags=default_unify_flags ()) (op,cl) = in ffail 0 in let rec matchrec cl = - let cl = strip_outer_cast evd (EConstr.of_constr cl) in + let cl = strip_outer_cast evd cl in + let cl = EConstr.of_constr cl in (bind - (if closed0 cl + (if closed0 evd cl then return (fun () -> w_typed_unify env evd CONV flags op cl,cl) else fail "Bound 1") - (match kind_of_term cl with + (match EConstr.kind evd cl with | App (f,args) -> let n = Array.length args in assert (n>0); @@ -1835,16 +1897,18 @@ let w_unify_to_subterm_all env evd ?(flags=default_unify_flags ()) (op,cl) = | _ -> res let w_unify_to_subterm_list env evd flags hdmeta oplist t = + let open EConstr in List.fold_right (fun op (evd,l) -> - let op = whd_meta evd (EConstr.of_constr op) in - if isMeta op then + let op = whd_meta evd op in + let op = EConstr.of_constr op in + if isMeta evd op then if flags.allow_K_in_toplevel_higher_order_unification then (evd,op::l) - else error_abstraction_over_meta env evd hdmeta (destMeta op) + else error_abstraction_over_meta env evd hdmeta (destMeta evd op) else let allow_K = flags.allow_K_in_toplevel_higher_order_unification in let flags = - if occur_meta_or_existential evd (EConstr.of_constr op) || !keyed_unification then + if occur_meta_or_existential evd op || !keyed_unification then (* This is up to delta for subterms w/o metas ... *) flags else @@ -1853,7 +1917,7 @@ let w_unify_to_subterm_list env evd flags hdmeta oplist t = unify pre-existing non frozen evars of the goal or of the pattern *) set_no_delta_flags flags in - let t' = (strip_outer_cast evd (EConstr.of_constr op),t) in + let t' = (EConstr.of_constr (strip_outer_cast evd op),t) in let (evd',cl) = try if is_keyed_unification () then @@ -1869,11 +1933,11 @@ let w_unify_to_subterm_list env evd flags hdmeta oplist t = (* w_unify_to_subterm does not go through evars, so the next step, which was already in <= 8.4, is needed at least for compatibility of rewrite *) - dependent evd (EConstr.of_constr op) (EConstr.of_constr t) -> (evd,op) + dependent evd op t -> (evd,op) in if not allow_K && (* ensure we found a different instance *) - List.exists (fun op -> Term.eq_constr op cl) l + List.exists (fun op -> EConstr.eq_constr evd' op cl) l then error_non_linear_unification env evd hdmeta cl else (evd',cl::l)) oplist @@ -1884,8 +1948,9 @@ let secondOrderAbstraction env evd flags typ (p, oplist) = let flags = { flags with core_unify_flags = flags.subterm_unify_flags } in let (evd',cllist) = w_unify_to_subterm_list env evd flags p oplist typ in let typp = Typing.meta_type evd' p in + let typp = EConstr.of_constr typp in let evd',(pred,predtyp) = abstract_list_all env evd' typp typ cllist in - let evd', b = infer_conv ~pb:CUMUL env evd' (EConstr.of_constr predtyp) (EConstr.of_constr typp) in + let evd', b = infer_conv ~pb:CUMUL env evd' predtyp typp in if not b then error_wrong_abstraction_type env evd' (Evd.meta_name evd p) pred typp predtyp; @@ -1902,7 +1967,8 @@ let secondOrderAbstraction env evd flags typ (p, oplist) = let secondOrderDependentAbstraction env evd flags typ (p, oplist) = let typp = Typing.meta_type evd p in - let evd, pred = abstract_list_all_with_dependencies env evd (EConstr.of_constr typp) typ (List.map EConstr.of_constr oplist) in + let evd, pred = abstract_list_all_with_dependencies env evd (EConstr.of_constr typp) typ oplist in + let pred = EConstr.of_constr pred in w_merge env false flags.merge_unify_flags (evd,[p,pred,(Conv,TypeProcessed)],[]) @@ -1910,16 +1976,15 @@ let secondOrderAbstractionAlgo dep = if dep then secondOrderDependentAbstraction else secondOrderAbstraction let w_unify2 env evd flags dep cv_pb ty1 ty2 = - let inj = EConstr.Unsafe.to_constr in - let c1, oplist1 = whd_nored_stack evd (EConstr.of_constr ty1) in - let c2, oplist2 = whd_nored_stack evd (EConstr.of_constr ty2) in + let c1, oplist1 = whd_nored_stack evd ty1 in + let c2, oplist2 = whd_nored_stack evd ty2 in match EConstr.kind evd c1, EConstr.kind evd c2 with | Meta p1, _ -> (* Find the predicate *) - secondOrderAbstractionAlgo dep env evd flags ty2 (p1, List.map inj oplist1) + secondOrderAbstractionAlgo dep env evd flags ty2 (p1, oplist1) | _, Meta p2 -> (* Find the predicate *) - secondOrderAbstractionAlgo dep env evd flags ty1 (p2, List.map inj oplist2) + secondOrderAbstractionAlgo dep env evd flags ty1 (p2, oplist2) | _ -> error "w_unify2" (* The unique unification algorithm works like this: If the pattern is @@ -1943,8 +2008,9 @@ let w_unify2 env evd flags dep cv_pb ty1 ty2 = convertible and first-order otherwise. But if failed if e.g. the type of Meta(1) had meta-variables in it. *) let w_unify env evd cv_pb ?(flags=default_unify_flags ()) ty1 ty2 = - let hd1,l1 = decompose_appvect (whd_nored evd (EConstr.of_constr ty1)) in - let hd2,l2 = decompose_appvect (whd_nored evd (EConstr.of_constr ty2)) in + let open EConstr in + let hd1,l1 = decompose_app_vect evd (EConstr.of_constr (whd_nored evd ty1)) in + let hd2,l2 = decompose_app_vect evd (EConstr.of_constr (whd_nored evd ty2)) in let is_empty1 = Array.is_empty l1 in let is_empty2 = Array.is_empty l2 in match kind_of_term hd1, not is_empty1, kind_of_term hd2, not is_empty2 with diff --git a/pretyping/unification.mli b/pretyping/unification.mli index 0ad882a9f..41dcb8ed3 100644 --- a/pretyping/unification.mli +++ b/pretyping/unification.mli @@ -7,6 +7,7 @@ (************************************************************************) open Term +open EConstr open Environ open Evd @@ -70,6 +71,8 @@ exception PatternNotFound type prefix_of_inductive_support_flag = bool +type pending_constr = Evd.pending * constr + type abstraction_request = | AbstractPattern of prefix_of_inductive_support_flag * (types -> bool) * Names.Name.t * pending_constr * Locus.clause * bool | AbstractExact of Names.Name.t * constr * types option * Locus.clause * bool @@ -97,28 +100,29 @@ val abstract_list_all : (* For tracing *) -val w_merge : env -> bool -> core_unify_flags -> evar_map * - (metavariable * constr * (instance_constraint * instance_typing_status)) list * - (env * types pexistential * types) list -> evar_map +type metabinding = (metavariable * constr * (instance_constraint * instance_typing_status)) + +type subst0 = + (evar_map * + metabinding list * + (Environ.env * existential * t) list) + +val w_merge : env -> bool -> core_unify_flags -> subst0 -> evar_map val unify_0 : Environ.env -> Evd.evar_map -> Evd.conv_pb -> core_unify_flags -> - Term.types -> - Term.types -> - Evd.evar_map * Evd.metabinding list * - (Environ.env * Term.types Term.pexistential * Term.constr) list + types -> + types -> + subst0 val unify_0_with_initial_metas : - Evd.evar_map * Evd.metabinding list * - (Environ.env * Term.types Term.pexistential * Term.constr) list -> + subst0 -> bool -> Environ.env -> Evd.conv_pb -> core_unify_flags -> - Term.types -> - Term.types -> - Evd.evar_map * Evd.metabinding list * - (Environ.env * Term.types Term.pexistential * Term.constr) list - + types -> + types -> + subst0 diff --git a/proofs/clenv.ml b/proofs/clenv.ml index 68aeaef5e..2d621e97c 100644 --- a/proofs/clenv.ml +++ b/proofs/clenv.ml @@ -257,7 +257,7 @@ let clenv_dependent ce = clenv_dependent_gen false ce let clenv_unify ?(flags=default_unify_flags ()) cv_pb t1 t2 clenv = { clenv with - evd = w_unify ~flags clenv.env clenv.evd cv_pb t1 t2 } + evd = w_unify ~flags clenv.env clenv.evd cv_pb (EConstr.of_constr t1) (EConstr.of_constr t2) } let clenv_unify_meta_types ?(flags=default_unify_flags ()) clenv = { clenv with evd = w_unify_meta_types ~flags:flags clenv.env clenv.evd } @@ -482,13 +482,15 @@ let error_already_defined b = (str "Position " ++ int n ++ str" already defined.") let clenv_unify_binding_type clenv c t u = - if isMeta (fst (decompose_appvect (whd_nored clenv.evd (EConstr.of_constr u)))) then + let u = EConstr.of_constr u in + if isMeta (fst (decompose_appvect (whd_nored clenv.evd u))) then (* Not enough information to know if some subtyping is needed *) CoerceToType, clenv, c else (* Enough information so as to try a coercion now *) try - let evd,c = w_coerce_to_type (cl_env clenv) clenv.evd c t u in + let evd,c = w_coerce_to_type (cl_env clenv) clenv.evd (EConstr.of_constr c) (EConstr.of_constr t) u in + let c = EConstr.Unsafe.to_constr c in TypeProcessed, { clenv with evd = evd }, c with | PretypeError (_,_,ActualTypeNotCoercible (_,_, diff --git a/proofs/clenvtac.ml b/proofs/clenvtac.ml index d8a20e08b..bc5f17bf5 100644 --- a/proofs/clenvtac.ml +++ b/proofs/clenvtac.ml @@ -143,7 +143,7 @@ let unify ?(flags=fail_quick_unif_flags) m = let n = Tacmach.New.pf_concl (Proofview.Goal.assume gl) in let evd = clear_metas (Tacmach.New.project gl) in try - let evd' = w_unify env evd CONV ~flags m n in + let evd' = w_unify env evd CONV ~flags (EConstr.of_constr m) (EConstr.of_constr n) in Proofview.Unsafe.tclEVARSADVANCE evd' with e when CErrors.noncritical e -> Proofview.tclZERO e end } diff --git a/proofs/evar_refiner.ml b/proofs/evar_refiner.ml index fd6528a1e..4be03af9a 100644 --- a/proofs/evar_refiner.ml +++ b/proofs/evar_refiner.ml @@ -26,7 +26,7 @@ let depends_on_evar evk _ (pbty,_,t1,t2) = let define_and_solve_constraints evk c env evd = if Termops.occur_evar evd evk (EConstr.of_constr c) then - Pretype_errors.error_occur_check env evd evk c; + Pretype_errors.error_occur_check env evd evk (EConstr.of_constr c); let evd = define evk c evd in let (evd,pbs) = extract_changed_conv_pbs evd (depends_on_evar evk) in match diff --git a/proofs/goal.ml b/proofs/goal.ml index a141708c2..bcb14e2d6 100644 --- a/proofs/goal.ml +++ b/proofs/goal.ml @@ -87,7 +87,7 @@ module V82 = struct (* Check that the goal itself does not appear in the refined term *) let _ = if not (Evarutil.occur_evar_upto sigma evk c) then () - else Pretype_errors.error_occur_check Environ.empty_env sigma evk c + else Pretype_errors.error_occur_check Environ.empty_env sigma evk (EConstr.of_constr c) in Evd.define evk c sigma diff --git a/proofs/refine.ml b/proofs/refine.ml index 19134bfa3..0ed74c9b3 100644 --- a/proofs/refine.ml +++ b/proofs/refine.ml @@ -98,7 +98,7 @@ let make_refine_enter ?(unsafe = true) f = let self = Proofview.Goal.goal gl in let _ = if not (Evarutil.occur_evar_upto sigma self c) then () - else Pretype_errors.error_occur_check env sigma self c + else Pretype_errors.error_occur_check env sigma self (EConstr.of_constr c) in (** Proceed to the refinement *) let sigma = match evkmain with diff --git a/tactics/equality.ml b/tactics/equality.ml index be175937b..64b56b99b 100644 --- a/tactics/equality.ml +++ b/tactics/equality.ml @@ -177,7 +177,7 @@ let instantiate_lemma_all frzevars gl c ty l l2r concl = let flags = make_flags frzevars (Tacmach.New.project gl) rewrite_unif_flags eqclause in let occs = w_unify_to_subterm_all ~flags env eqclause.evd - ((if l2r then c1 else c2),concl) + (EConstr.of_constr (if l2r then c1 else c2),EConstr.of_constr concl) in List.map try_occ occs let instantiate_lemma gl c ty l l2r concl = diff --git a/tactics/inv.ml b/tactics/inv.ml index 9282af759..9324d8e37 100644 --- a/tactics/inv.ml +++ b/tactics/inv.ml @@ -90,8 +90,8 @@ let make_inv_predicate env evd indf realargs id status concl = let sort = Evarutil.evd_comb1 (Evd.fresh_sort_in_family env) evd sort in let p = make_arity env true indf sort in let evd',(p,ptyp) = Unification.abstract_list_all env - !evd p concl (realargs@[mkVar id]) - in evd := evd'; p in + !evd (EConstr.of_constr p) (EConstr.of_constr concl) (List.map EConstr.of_constr realargs@[EConstr.mkVar id]) + in evd := evd'; EConstr.Unsafe.to_constr p in let hyps,bodypred = decompose_lam_n_assum (nrealargs+1) pred in (* We lift to make room for the equations *) (hyps,lift nrealargs bodypred) diff --git a/tactics/tactics.ml b/tactics/tactics.ml index 3bb285aa8..2cb9e0864 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -1204,10 +1204,12 @@ let map_destruction_arg f sigma = function let finish_delayed_evar_resolution with_evars env sigma f = let ((c, lbind), sigma') = run_delayed env sigma f in + let c = EConstr.of_constr c in let pending = (sigma,sigma') in let sigma' = Sigma.Unsafe.of_evar_map sigma' in let flags = tactic_infer_flags with_evars in let Sigma (c, sigma', _) = finish_evar_resolution ~flags env sigma' (pending,c) in + let c = EConstr.Unsafe.to_constr c in (Sigma.to_evar_map sigma', (c, lbind)) let with_no_bindings (c, lbind) = @@ -2692,14 +2694,18 @@ let letin_tac with_eq id c ty occs = let sigma = Proofview.Goal.sigma gl in let env = Proofview.Goal.env gl in let ccl = Proofview.Goal.concl gl in - let abs = AbstractExact (id,c,ty,occs,true) in + let c = EConstr.of_constr c in + let abs = AbstractExact (id,c,Option.map EConstr.of_constr ty,occs,true) in + let ccl = EConstr.of_constr ccl in let (id,_,depdecls,lastlhyp,ccl,res) = make_abstraction env sigma ccl abs in + let ccl = EConstr.Unsafe.to_constr ccl in (* We keep the original term to match but record the potential side-effects of unifying universes. *) let Sigma (c, sigma, p) = match res with | None -> Sigma.here c sigma | Some (Sigma (_, sigma, p)) -> Sigma (c, sigma, p) in + let c = EConstr.Unsafe.to_constr c in let tac = letin_tac_gen with_eq (id,depdecls,lastlhyp,ccl,c) ty in Sigma (tac, sigma, p) end } @@ -2711,10 +2717,13 @@ let letin_pat_tac with_eq id c occs = let ccl = Proofview.Goal.concl gl in let check t = true in let abs = AbstractPattern (false,check,id,c,occs,false) in + let ccl = EConstr.of_constr ccl in let (id,_,depdecls,lastlhyp,ccl,res) = make_abstraction env sigma ccl abs in + let ccl = EConstr.Unsafe.to_constr ccl in let Sigma (c, sigma, p) = match res with | None -> finish_evar_resolution ~flags:(tactic_infer_flags false) env sigma c | Some res -> res in + let c = EConstr.Unsafe.to_constr c in let tac = (letin_tac_gen with_eq (id,depdecls,lastlhyp,ccl,c) None) in @@ -4263,8 +4272,8 @@ let use_bindings env sigma elim must_be_closed (c,lbind) typ = if must_be_closed && occur_meta indclause.evd (EConstr.of_constr (clenv_value indclause)) then error "Need a fully applied argument."; (* We lose the possibility of coercions in with-bindings *) - let (sigma, c) = pose_all_metas_as_evars env indclause.evd (clenv_value indclause) in - Sigma.Unsafe.of_pair (c, sigma) + let (sigma, c) = pose_all_metas_as_evars env indclause.evd (EConstr.of_constr (clenv_value indclause)) in + Sigma.Unsafe.of_pair (EConstr.Unsafe.to_constr c, sigma) with e when catchable_exception e -> try find_clause (try_red_product env sigma (EConstr.of_constr typ)) with Redelimination -> raise e in @@ -4279,7 +4288,7 @@ let check_expected_type env sigma (elimc,bl) elimt = let sigma,cl = make_evar_clause env sigma ~len:(n - 1) elimt in let sigma = solve_evar_clause env sigma true cl bl in let (_,u,_) = destProd cl.cl_concl in - fun t -> Evarconv.e_cumul env (ref sigma) (EConstr.of_constr t) (EConstr.of_constr u) + fun t -> Evarconv.e_cumul env (ref sigma) t (EConstr.of_constr u) let check_enough_applied env sigma elim = let sigma = Sigma.to_evar_map sigma in @@ -4288,7 +4297,7 @@ let check_enough_applied env sigma elim = | None -> (* No eliminator given *) fun u -> - let t,_ = decompose_app (whd_all env sigma (EConstr.of_constr u)) in isInd t + let t,_ = decompose_app (whd_all env sigma u) in isInd t | Some elimc -> let elimt = Retyping.get_type_of env sigma (EConstr.of_constr (fst elimc)) in let scheme = compute_elim_sig ~elimc elimt in @@ -4314,8 +4323,11 @@ let pose_induction_arg_then isrec with_evars (is_arg_pure_hyp,from_prefix) elim let store = Proofview.Goal.extra gl in let check = check_enough_applied env sigma elim in let Sigma (c, sigma', p) = use_bindings env sigma elim false (c0,lbind) t0 in + let c = EConstr.of_constr c in let abs = AbstractPattern (from_prefix,check,Name id,(pending,c),cls,false) in + let ccl = EConstr.of_constr ccl in let (id,sign,_,lastlhyp,ccl,res) = make_abstraction env sigma' ccl abs in + let ccl = EConstr.Unsafe.to_constr ccl in match res with | None -> (* pattern not found *) @@ -4323,7 +4335,9 @@ let pose_induction_arg_then isrec with_evars (is_arg_pure_hyp,from_prefix) elim (* we restart using bindings after having tried type-class resolution etc. on the term given by the user *) let flags = tactic_infer_flags (with_evars && (* do not give a success semantics to edestruct on an open term yet *) false) in + let c0 = EConstr.of_constr c0 in let Sigma (c0, sigma, q) = finish_evar_resolution ~flags env sigma (pending,c0) in + let c0 = EConstr.Unsafe.to_constr c0 in let tac = (if isrec then (* Historically, induction has side conditions last *) @@ -4350,6 +4364,7 @@ let pose_induction_arg_then isrec with_evars (is_arg_pure_hyp,from_prefix) elim Sigma (tac, sigma, q) | Some (Sigma (c, sigma', q)) -> + let c = EConstr.Unsafe.to_constr c in (* pattern found *) let with_eq = Option.map (fun eq -> (false,eq)) eqname in (* TODO: if ind has predicate parameters, use JMeq instead of eq *) @@ -4386,7 +4401,7 @@ let induction_gen clear_flag isrec with_evars elim && lbind == NoBindings && not with_evars && Option.is_empty eqname && clear_flag == None && has_generic_occurrences_but_goal cls (destVar c) env (Sigma.to_evar_map sigma) ccl in - let enough_applied = check_enough_applied env sigma elim t in + let enough_applied = check_enough_applied env sigma elim (EConstr.of_constr t) in if is_arg_pure_hyp && enough_applied then (* First case: induction on a variable already in an inductive type and with maximal abstraction over the variable. @@ -4935,6 +4950,8 @@ let tclABSTRACT name_op tac = abstract_subproof s gk tac let unify ?(state=full_transparent_state) x y = + let x = EConstr.of_constr x in + let y = EConstr.of_constr y in Proofview.Goal.nf_s_enter { s_enter = begin fun gl -> let sigma = Proofview.Goal.sigma gl in try diff --git a/toplevel/himsg.ml b/toplevel/himsg.ml index 345eae9df..20cda947a 100644 --- a/toplevel/himsg.ml +++ b/toplevel/himsg.ml @@ -272,7 +272,7 @@ let explain_ill_formed_branch env sigma c ci actty expty = let explain_generalization env sigma (name,var) j = let pe = pr_ne_context_of (str "In environment") env sigma in let pv = pr_ltype_env env sigma var in - let (pc,pt) = pr_ljudge_env (push_rel_assum (name,var) env) sigma j in + let (pc,pt) = pr_ljudge_env (push_rel_assum (name,EConstr.of_constr var) env) sigma j in pe ++ str "Cannot generalize" ++ brk(1,1) ++ pv ++ spc () ++ str "over" ++ brk(1,1) ++ pc ++ str "," ++ spc () ++ str "it has type" ++ spc () ++ pt ++ @@ -529,7 +529,7 @@ let explain_cant_find_case_type env sigma c = pe ++ str "." let explain_occur_check env sigma ev rhs = - let rhs = Evarutil.nf_evar sigma rhs in + let rhs = EConstr.to_constr sigma rhs in let env = make_all_name_different env in let pt = pr_lconstr_env env sigma rhs in str "Cannot define " ++ pr_existential_key sigma ev ++ str " with term" ++ @@ -637,6 +637,9 @@ let explain_cannot_unify env sigma m n e = str "with" ++ brk(1,1) ++ pn ++ ppreason ++ str "." let explain_cannot_unify_local env sigma m n subn = + let m = EConstr.to_constr sigma m in + let n = EConstr.to_constr sigma n in + let subn = EConstr.to_constr sigma subn in let pm = pr_lconstr_env env sigma m in let pn = pr_lconstr_env env sigma n in let psubn = pr_lconstr_env env sigma subn in @@ -649,6 +652,7 @@ let explain_refiner_cannot_generalize env sigma ty = pr_lconstr_env env sigma ty ++ str "." let explain_no_occurrence_found env sigma c id = + let c = EConstr.to_constr sigma c in str "Found no subterm matching " ++ pr_lconstr_env env sigma c ++ str " in " ++ (match id with @@ -662,6 +666,7 @@ let explain_cannot_unify_binding_type env sigma m n = str "which should be unifiable with" ++ brk(1,1) ++ pn ++ str "." let explain_cannot_find_well_typed_abstraction env sigma p l e = + let p = EConstr.to_constr sigma p in str "Abstracting over the " ++ str (String.plural (List.length l) "term") ++ spc () ++ hov 0 (pr_enum (fun c -> pr_lconstr_env env sigma (EConstr.to_constr sigma c)) l) ++ spc () ++ @@ -670,6 +675,9 @@ let explain_cannot_find_well_typed_abstraction env sigma p l e = (match e with None -> mt () | Some e -> fnl () ++ str "Reason is: " ++ e) let explain_wrong_abstraction_type env sigma na abs expected result = + let abs = EConstr.to_constr sigma abs in + let expected = EConstr.to_constr sigma expected in + let result = EConstr.to_constr sigma result in let ppname = match na with Name id -> pr_id id ++ spc () | _ -> mt () in str "Cannot instantiate metavariable " ++ ppname ++ strbrk "of type " ++ pr_lconstr_env env sigma expected ++ strbrk " with abstraction " ++ @@ -681,6 +689,7 @@ let explain_abstraction_over_meta _ m n = pr_name m ++ spc () ++ str "and " ++ pr_name n ++ str "." let explain_non_linear_unification env sigma m t = + let t = EConstr.to_constr sigma t in strbrk "Cannot unambiguously instantiate " ++ pr_name m ++ str ":" ++ strbrk " which would require to abstract twice on " ++ @@ -742,7 +751,13 @@ let explain_cannot_unify_occurrences env sigma nested ((cl2,pos2),t2) ((cl1,pos1 str "Found nested occurrences of the pattern at positions " ++ int pos1 ++ strbrk " and " ++ pr_position (cl2,pos2) ++ str "." else - let ppreason = match e with None -> mt() | Some (c1,c2,e) -> explain_unification_error env sigma c1 c2 (Some e) in + let ppreason = match e with + | None -> mt() + | Some (c1,c2,e) -> + let c1 = EConstr.to_constr sigma c1 in + let c2 = EConstr.to_constr sigma c2 in + explain_unification_error env sigma c1 c2 (Some e) + in str "Found incompatible occurrences of the pattern" ++ str ":" ++ spc () ++ str "Matched term " ++ pr_lconstr_env env sigma (EConstr.to_constr sigma t2) ++ strbrk " at position " ++ pr_position (cl2,pos2) ++ @@ -819,6 +834,8 @@ let explain_pretype_error env sigma err = explain_unexpected_type env sigma actual expect | NotProduct c -> explain_not_product env sigma c | CannotUnify (m,n,e) -> + let m = EConstr.Unsafe.to_constr m in + let n = EConstr.Unsafe.to_constr n in let env, m, n = contract2 env m n in explain_cannot_unify env sigma m n e | CannotUnifyLocal (m,n,sn) -> explain_cannot_unify_local env sigma m n sn -- cgit v1.2.3 From ca993b9e7765ac58f70740818758457c9367b0da Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 11 Nov 2016 00:29:02 +0100 Subject: Making judgment type generic over the type of inner constrs. This allows to factorize code and prevents the unnecessary use of back and forth conversions between the various types of terms. Note that functions from typing may now raise errors as PretypeError rather than TypeError, because they call the proper wrapper. I think that they were wrongly calling the kernel because of an overlook of open modules. --- engine/eConstr.ml | 2 + engine/eConstr.mli | 2 + engine/evarutil.ml | 7 +- engine/evarutil.mli | 8 +- engine/termops.ml | 1 + engine/termops.mli | 8 +- kernel/environ.ml | 14 ++-- kernel/environ.mli | 19 +++-- kernel/type_errors.ml | 58 +++++++------ kernel/type_errors.mli | 58 +++++++------ ltac/tacinterp.ml | 4 +- plugins/cc/cctac.ml | 2 +- pretyping/cases.ml | 57 ++++++------- pretyping/cases.mli | 6 +- pretyping/classops.ml | 2 +- pretyping/classops.mli | 2 +- pretyping/coercion.ml | 60 ++++++++------ pretyping/coercion.mli | 9 +- pretyping/evarconv.ml | 4 +- pretyping/pretype_errors.ml | 24 ++++-- pretyping/pretype_errors.mli | 41 ++++++---- pretyping/pretyping.ml | 157 ++++++++++++++++++----------------- pretyping/pretyping.mli | 12 +-- pretyping/retyping.ml | 2 +- pretyping/retyping.mli | 2 +- pretyping/typing.ml | 191 +++++++++++++++++++++++++++++-------------- pretyping/typing.mli | 10 ++- pretyping/unification.ml | 8 +- proofs/clenv.ml | 6 +- proofs/refine.ml | 5 +- proofs/refine.mli | 2 +- tactics/contradiction.ml | 2 +- toplevel/explainErr.ml | 1 + toplevel/himsg.ml | 72 +++++++++++++++- toplevel/himsg.mli | 3 + toplevel/vernacentries.ml | 3 +- 36 files changed, 538 insertions(+), 326 deletions(-) diff --git a/engine/eConstr.ml b/engine/eConstr.ml index 0a5f1ba68..095521e25 100644 --- a/engine/eConstr.ml +++ b/engine/eConstr.ml @@ -75,6 +75,8 @@ 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 let mkProp = of_kind (Sort Sorts.prop) let mkSet = of_kind (Sort Sorts.set) diff --git a/engine/eConstr.mli b/engine/eConstr.mli index 03e2d4ffc..10eb064a3 100644 --- a/engine/eConstr.mli +++ b/engine/eConstr.mli @@ -20,6 +20,8 @@ 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 (** {5 Destructors} *) diff --git a/engine/evarutil.ml b/engine/evarutil.ml index fc193b94a..bc55ac458 100644 --- a/engine/evarutil.ml +++ b/engine/evarutil.ml @@ -96,14 +96,15 @@ let rec whd_evar sigma c = | _ -> c let rec nf_evar sigma t = Constr.map (fun t -> nf_evar sigma t) (whd_evar sigma t) +let e_nf_evar sigma t = EConstr.of_constr (nf_evar sigma (EConstr.Unsafe.to_constr t)) let j_nf_evar sigma j = - { uj_val = nf_evar sigma j.uj_val; - uj_type = nf_evar sigma j.uj_type } + { uj_val = e_nf_evar sigma j.uj_val; + uj_type = e_nf_evar sigma j.uj_type } let jl_nf_evar sigma jl = List.map (j_nf_evar sigma) jl let jv_nf_evar sigma = Array.map (j_nf_evar sigma) let tj_nf_evar sigma {utj_val=v;utj_type=t} = - {utj_val=nf_evar sigma v;utj_type=t} + {utj_val=e_nf_evar sigma v;utj_type=t} let nf_evars_universes evm = Universes.nf_evars_and_universes_opt_subst (safe_evar_value evm) diff --git a/engine/evarutil.mli b/engine/evarutil.mli index dcb9bf247..8f3c3c352 100644 --- a/engine/evarutil.mli +++ b/engine/evarutil.mli @@ -141,13 +141,13 @@ val judge_of_new_Type : 'r Sigma.t -> (unsafe_judgment, 'r) Sigma.sigma val whd_evar : evar_map -> constr -> constr val nf_evar : evar_map -> constr -> constr -val j_nf_evar : evar_map -> unsafe_judgment -> unsafe_judgment +val j_nf_evar : evar_map -> EConstr.unsafe_judgment -> EConstr.unsafe_judgment val jl_nf_evar : - evar_map -> unsafe_judgment list -> unsafe_judgment list + evar_map -> EConstr.unsafe_judgment list -> EConstr.unsafe_judgment list val jv_nf_evar : - evar_map -> unsafe_judgment array -> unsafe_judgment array + evar_map -> EConstr.unsafe_judgment array -> EConstr.unsafe_judgment array val tj_nf_evar : - evar_map -> unsafe_type_judgment -> unsafe_type_judgment + evar_map -> EConstr.unsafe_type_judgment -> EConstr.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 diff --git a/engine/termops.ml b/engine/termops.ml index c1198e05a..83f07d2c6 100644 --- a/engine/termops.ml +++ b/engine/termops.ml @@ -1141,6 +1141,7 @@ let impossible_default_case = ref None let set_impossible_default_clause c = impossible_default_case := Some c let coq_unit_judge = + let make_judge c t = make_judge (EConstr.of_constr c) (EConstr.of_constr t) in let na1 = Name (Id.of_string "A") in let na2 = Name (Id.of_string "H") in fun () -> diff --git a/engine/termops.mli b/engine/termops.mli index df3fdd124..27b3ea53c 100644 --- a/engine/termops.mli +++ b/engine/termops.mli @@ -274,10 +274,10 @@ val is_template_polymorphic : env -> Evd.evar_map -> EConstr.t -> 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 (** {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 +val coq_unit_judge : unit -> EConstr.unsafe_judgment Univ.in_universe_context_set diff --git a/kernel/environ.ml b/kernel/environ.ml index 4a543f195..5688813ad 100644 --- a/kernel/environ.ml +++ b/kernel/environ.ml @@ -469,9 +469,11 @@ let lookup_modtype mp env = (*s Judgments. *) -type unsafe_judgment = { - uj_val : constr; - uj_type : types } +type ('constr, 'types) punsafe_judgment = { + uj_val : 'constr; + uj_type : 'types } + +type unsafe_judgment = (constr, types) punsafe_judgment let make_judge v tj = { uj_val = v; @@ -480,10 +482,12 @@ let make_judge v tj = let j_val j = j.uj_val let j_type j = j.uj_type -type unsafe_type_judgment = { - utj_val : constr; +type 'types punsafe_type_judgment = { + utj_val : 'types; utj_type : sorts } +type unsafe_type_judgment = types punsafe_type_judgment + (*s Compilation of global declaration *) let compile_constant_body = Cbytegen.compile_constant_body false diff --git a/kernel/environ.mli b/kernel/environ.mli index ea570cb4a..b7431dbe5 100644 --- a/kernel/environ.mli +++ b/kernel/environ.mli @@ -238,18 +238,21 @@ val keep_hyps : env -> Id.Set.t -> Context.Named.t actually only a datatype to store a term with its type and the type of its type. *) -type unsafe_judgment = { - uj_val : constr; - uj_type : types } +type ('constr, 'types) punsafe_judgment = { + uj_val : 'constr; + uj_type : 'types } -val make_judge : constr -> types -> unsafe_judgment -val j_val : unsafe_judgment -> constr -val j_type : unsafe_judgment -> types +type unsafe_judgment = (constr, types) punsafe_judgment -type unsafe_type_judgment = { - utj_val : constr; +val make_judge : 'constr -> 'types -> ('constr, 'types) punsafe_judgment +val j_val : ('constr, 'types) punsafe_judgment -> 'constr +val j_type : ('constr, 'types) punsafe_judgment -> 'types + +type 'types punsafe_type_judgment = { + utj_val : 'types; utj_type : sorts } +type unsafe_type_judgment = types punsafe_type_judgment (** {6 Compilation of global declaration } *) diff --git a/kernel/type_errors.ml b/kernel/type_errors.ml index 5071f0ad5..5e1763815 100644 --- a/kernel/type_errors.ml +++ b/kernel/type_errors.ml @@ -13,52 +13,56 @@ open Reduction (* Type errors. *) -type guard_error = +type 'constr pguard_error = (* Fixpoints *) | NotEnoughAbstractionInFixBody - | RecursionNotOnInductiveType of constr - | RecursionOnIllegalTerm of int * (env * constr) * int list * int list + | RecursionNotOnInductiveType of 'constr + | RecursionOnIllegalTerm of int * (env * 'constr) * int list * int list | NotEnoughArgumentsForFixCall of int (* CoFixpoints *) - | CodomainNotInductiveType of constr + | CodomainNotInductiveType of 'constr | NestedRecursiveOccurrences - | UnguardedRecursiveCall of constr - | RecCallInTypeOfAbstraction of constr - | RecCallInNonRecArgOfConstructor of constr - | RecCallInTypeOfDef of constr - | RecCallInCaseFun of constr - | RecCallInCaseArg of constr - | RecCallInCasePred of constr - | NotGuardedForm of constr - | ReturnPredicateNotCoInductive of constr + | UnguardedRecursiveCall of 'constr + | RecCallInTypeOfAbstraction of 'constr + | RecCallInNonRecArgOfConstructor of 'constr + | RecCallInTypeOfDef of 'constr + | RecCallInCaseFun of 'constr + | RecCallInCaseArg of 'constr + | RecCallInCasePred of 'constr + | NotGuardedForm of 'constr + | ReturnPredicateNotCoInductive of 'constr + +type guard_error = constr pguard_error type arity_error = | NonInformativeToInformative | StrongEliminationOnNonSmallType | WrongArity -type type_error = +type ('constr, 'types) ptype_error = | UnboundRel of int | UnboundVar of variable - | NotAType of unsafe_judgment - | BadAssumption of unsafe_judgment - | ReferenceVariables of identifier * constr - | ElimArity of pinductive * sorts_family list * constr * unsafe_judgment + | NotAType of ('constr, 'types) punsafe_judgment + | BadAssumption of ('constr, 'types) punsafe_judgment + | ReferenceVariables of identifier * 'constr + | ElimArity of pinductive * sorts_family list * 'constr * ('constr, 'types) punsafe_judgment * (sorts_family * sorts_family * arity_error) option - | CaseNotInductive of unsafe_judgment + | CaseNotInductive of ('constr, 'types) punsafe_judgment | WrongCaseInfo of pinductive * case_info - | NumberBranches of unsafe_judgment * int - | IllFormedBranch of constr * pconstructor * constr * constr - | Generalization of (Name.t * types) * unsafe_judgment - | ActualType of unsafe_judgment * types + | NumberBranches of ('constr, 'types) punsafe_judgment * int + | IllFormedBranch of 'constr * pconstructor * 'constr * 'constr + | Generalization of (Name.t * 'types) * ('constr, 'types) punsafe_judgment + | ActualType of ('constr, 'types) punsafe_judgment * 'types | CantApplyBadType of - (int * constr * constr) * unsafe_judgment * unsafe_judgment array - | CantApplyNonFunctional of unsafe_judgment * unsafe_judgment array - | IllFormedRecBody of guard_error * Name.t array * int * env * unsafe_judgment array + (int * 'constr * 'constr) * ('constr, 'types) punsafe_judgment * ('constr, 'types) punsafe_judgment array + | CantApplyNonFunctional of ('constr, 'types) punsafe_judgment * ('constr, 'types) punsafe_judgment array + | IllFormedRecBody of 'constr pguard_error * Name.t array * int * env * ('constr, 'types) punsafe_judgment array | IllTypedRecBody of - int * Name.t array * unsafe_judgment array * types array + int * Name.t array * ('constr, 'types) punsafe_judgment array * 'types array | UnsatisfiedConstraints of Univ.constraints +type type_error = (constr, types) ptype_error + exception TypeError of env * type_error let nfj env {uj_val=c;uj_type=ct} = diff --git a/kernel/type_errors.mli b/kernel/type_errors.mli index 0c3a952b8..bd6032716 100644 --- a/kernel/type_errors.mli +++ b/kernel/type_errors.mli @@ -14,52 +14,56 @@ open Environ (*i Rem: NotEnoughAbstractionInFixBody should only occur with "/i" Fix notation i*) -type guard_error = +type 'constr pguard_error = (** Fixpoints *) | NotEnoughAbstractionInFixBody - | RecursionNotOnInductiveType of constr - | RecursionOnIllegalTerm of int * (env * constr) * int list * int list + | RecursionNotOnInductiveType of 'constr + | RecursionOnIllegalTerm of int * (env * 'constr) * int list * int list | NotEnoughArgumentsForFixCall of int (** CoFixpoints *) - | CodomainNotInductiveType of constr + | CodomainNotInductiveType of 'constr | NestedRecursiveOccurrences - | UnguardedRecursiveCall of constr - | RecCallInTypeOfAbstraction of constr - | RecCallInNonRecArgOfConstructor of constr - | RecCallInTypeOfDef of constr - | RecCallInCaseFun of constr - | RecCallInCaseArg of constr - | RecCallInCasePred of constr - | NotGuardedForm of constr - | ReturnPredicateNotCoInductive of constr + | UnguardedRecursiveCall of 'constr + | RecCallInTypeOfAbstraction of 'constr + | RecCallInNonRecArgOfConstructor of 'constr + | RecCallInTypeOfDef of 'constr + | RecCallInCaseFun of 'constr + | RecCallInCaseArg of 'constr + | RecCallInCasePred of 'constr + | NotGuardedForm of 'constr + | ReturnPredicateNotCoInductive of 'constr + +type guard_error = constr pguard_error type arity_error = | NonInformativeToInformative | StrongEliminationOnNonSmallType | WrongArity -type type_error = +type ('constr, 'types) ptype_error = | UnboundRel of int | UnboundVar of variable - | NotAType of unsafe_judgment - | BadAssumption of unsafe_judgment - | ReferenceVariables of identifier * constr - | ElimArity of pinductive * sorts_family list * constr * unsafe_judgment + | NotAType of ('constr, 'types) punsafe_judgment + | BadAssumption of ('constr, 'types) punsafe_judgment + | ReferenceVariables of identifier * 'constr + | ElimArity of pinductive * sorts_family list * 'constr * ('constr, 'types) punsafe_judgment * (sorts_family * sorts_family * arity_error) option - | CaseNotInductive of unsafe_judgment + | CaseNotInductive of ('constr, 'types) punsafe_judgment | WrongCaseInfo of pinductive * case_info - | NumberBranches of unsafe_judgment * int - | IllFormedBranch of constr * pconstructor * constr * constr - | Generalization of (Name.t * types) * unsafe_judgment - | ActualType of unsafe_judgment * types + | NumberBranches of ('constr, 'types) punsafe_judgment * int + | IllFormedBranch of 'constr * pconstructor * 'constr * 'constr + | Generalization of (Name.t * 'types) * ('constr, 'types) punsafe_judgment + | ActualType of ('constr, 'types) punsafe_judgment * 'types | CantApplyBadType of - (int * constr * constr) * unsafe_judgment * unsafe_judgment array - | CantApplyNonFunctional of unsafe_judgment * unsafe_judgment array - | IllFormedRecBody of guard_error * Name.t array * int * env * unsafe_judgment array + (int * 'constr * 'constr) * ('constr, 'types) punsafe_judgment * ('constr, 'types) punsafe_judgment array + | CantApplyNonFunctional of ('constr, 'types) punsafe_judgment * ('constr, 'types) punsafe_judgment array + | IllFormedRecBody of 'constr pguard_error * Name.t array * int * env * ('constr, 'types) punsafe_judgment array | IllTypedRecBody of - int * Name.t array * unsafe_judgment array * types array + int * Name.t array * ('constr, 'types) punsafe_judgment array * 'types array | UnsatisfiedConstraints of Univ.constraints +type type_error = (constr, types) ptype_error + exception TypeError of env * type_error val error_unbound_rel : env -> int -> 'a diff --git a/ltac/tacinterp.ml b/ltac/tacinterp.ml index d8c933912..142f061b5 100644 --- a/ltac/tacinterp.ml +++ b/ltac/tacinterp.ml @@ -2100,11 +2100,13 @@ let interp_redexp env sigma r = let _ = let eval ty env sigma lfun arg = + let ty = EConstr.Unsafe.to_constr ty in let ist = { lfun = lfun; extra = TacStore.empty; } in if Genarg.has_type arg (glbwit wit_tactic) then let tac = Genarg.out_gen (glbwit wit_tactic) arg in let tac = interp_tactic ist tac in - Pfedit.refine_by_tactic env sigma ty tac + let (c, sigma) = Pfedit.refine_by_tactic env sigma ty tac in + (EConstr.of_constr c, sigma) else failwith "not a tactic" in diff --git a/plugins/cc/cctac.ml b/plugins/cc/cctac.ml index 58454eedf..11da923e1 100644 --- a/plugins/cc/cctac.ml +++ b/plugins/cc/cctac.ml @@ -512,7 +512,7 @@ let f_equal = | _ -> Proofview.tclUNIT () end begin function (e, info) -> match e with - | Type_errors.TypeError _ -> Proofview.tclUNIT () + | Pretype_errors.PretypeError _ | Type_errors.TypeError _ -> Proofview.tclUNIT () | e -> Proofview.tclZERO ~info e end end } diff --git a/pretyping/cases.ml b/pretyping/cases.ml index 92bd1e389..b43e2193a 100644 --- a/pretyping/cases.ml +++ b/pretyping/cases.ml @@ -89,9 +89,6 @@ let list_try_compile f l = let force_name = let nx = Name default_dependent_ident in function Anonymous -> nx | na -> na -let make_judge c ty = - make_judge (EConstr.Unsafe.to_constr c) (EConstr.Unsafe.to_constr ty) - (************************************************************************) (* Pattern-matching compilation (Cases) *) (************************************************************************) @@ -265,7 +262,7 @@ type 'a pattern_matching_problem = mat : 'a matrix; caseloc : Loc.t; casestyle : case_style; - typing_function: type_constraint -> env -> evar_map ref -> 'a option -> unsafe_judgment } + typing_function: type_constraint -> env -> evar_map ref -> 'a option -> EConstr.unsafe_judgment } (*--------------------------------------------------------------------------* * A few functions to infer the inductive type from the patterns instead of * @@ -366,12 +363,12 @@ let coerce_row typing_fun evdref env pats (tomatch,(_,indopt)) = let j = typing_fun tycon env evdref tomatch in let evd, j = Coercion.inh_coerce_to_base (loc_of_glob_constr tomatch) env !evdref j in evdref := evd; - let typ = EConstr.of_constr (nf_evar !evdref j.uj_type) in + let typ = EConstr.of_constr (nf_evar !evdref (EConstr.Unsafe.to_constr j.uj_type)) in let t = try try_find_ind env !evdref typ realnames with Not_found -> unify_tomatch_with_patterns evdref env loc typ pats realnames in - (EConstr.of_constr j.uj_val,t) + (j.uj_val,t) let coerce_to_indtype typing_fun evdref env matx tomatchl = let pats = List.map (fun r -> r.patterns) matx in @@ -415,7 +412,7 @@ let adjust_tomatch_to_pattern pb ((current,typ),deps,dep) = let _ = e_cumul pb.env pb.evdref indt typ in current else - EConstr.of_constr (evd_comb2 (Coercion.inh_conv_coerce_to true Loc.ghost pb.env) + (evd_comb2 (Coercion.inh_conv_coerce_to true Loc.ghost pb.env) pb.evdref (make_judge current typ) indt).uj_val in let sigma = !(pb.evdref) in (current,try_find_ind pb.env sigma indt names)) @@ -1002,7 +999,7 @@ let adjust_impossible_cases pb pred tomatch submat = | Evar (evk,_) when snd (evar_source evk !(pb.evdref)) == Evar_kinds.ImpossibleCase -> if not (Evd.is_defined !(pb.evdref) evk) then begin let evd, default = use_unit_judge !(pb.evdref) in - pb.evdref := Evd.define evk default.uj_type evd + pb.evdref := Evd.define evk (EConstr.Unsafe.to_constr default.uj_type) evd end; add_assert_false_case pb tomatch | _ -> @@ -1411,8 +1408,8 @@ and match_current pb (initial,tomatch) = make_case_or_project pb.env !(pb.evdref) indf ci pred current brvals in Typing.check_allowed_sort pb.env !(pb.evdref) mind current pred; - { uj_val = EConstr.Unsafe.to_constr (applist (case, inst)); - uj_type = EConstr.Unsafe.to_constr (prod_applist !(pb.evdref) typ inst) } + { uj_val = applist (case, inst); + uj_type = prod_applist !(pb.evdref) typ inst } (* Building the sub-problem when all patterns are variables. Case @@ -1429,8 +1426,8 @@ and shift_problem ((current,t),_,na) pb = history = pop_history pb.history; mat = List.map (push_current_pattern (current,ty)) pb.mat } in let j = compile pb in - { uj_val = EConstr.Unsafe.to_constr (subst1 current (EConstr.of_constr j.uj_val)); - uj_type = EConstr.Unsafe.to_constr (subst1 current (EConstr.of_constr j.uj_type)) } + { uj_val = subst1 current j.uj_val; + uj_type = subst1 current j.uj_type } (* Building the sub-problem when all patterns are variables, non-initial case. Variables which appear as subterms of constructor @@ -1453,7 +1450,7 @@ and compile_all_variables initial cur pb = (* Building the sub-problem when all patterns are variables *) and compile_branch initial current realargs names deps pb arsign eqns cstr = let sign, pb = build_branch initial current realargs deps names pb arsign eqns cstr in - sign, EConstr.of_constr (compile pb).uj_val + sign, (compile pb).uj_val (* Abstract over a declaration before continuing splitting *) and compile_generalization pb i d rest = @@ -1463,8 +1460,8 @@ and compile_generalization pb i d rest = tomatch = rest; mat = List.map (push_generalized_decl_eqn pb.env i d) pb.mat } in let j = compile pb in - { uj_val = Term.mkLambda_or_LetIn d j.uj_val; - uj_type = Term.mkProd_wo_LetIn d j.uj_type } + { uj_val = mkLambda_or_LetIn d j.uj_val; + uj_type = mkProd_wo_LetIn d j.uj_type } (* spiwack: the [initial] argument keeps track whether the alias has been introduced by a toplevel branch ([true]) or a deep one @@ -1482,11 +1479,11 @@ and compile_alias initial pb (na,orig,(expanded,expanded_typ)) rest = let j = compile pb in let sigma = !(pb.evdref) in { uj_val = - if isRel sigma c || isVar sigma c || count_occurrences sigma (mkRel 1) (EConstr.of_constr j.uj_val) <= 1 then - EConstr.Unsafe.to_constr (subst1 c (EConstr.of_constr j.uj_val)) + if isRel sigma c || isVar sigma c || count_occurrences sigma (mkRel 1) j.uj_val <= 1 then + subst1 c j.uj_val else - EConstr.Unsafe.to_constr (mkLetIn (na,c,t,EConstr.of_constr j.uj_val)); - uj_type = EConstr.Unsafe.to_constr (subst1 c (EConstr.of_constr j.uj_type)) } in + mkLetIn (na,c,t,j.uj_val); + uj_type = subst1 c j.uj_type } in (* spiwack: when an alias appears on a deep branch, its non-expanded form is automatically a variable of the same name. We avoid introducing such superfluous aliases so that refines are elegant. *) @@ -1726,7 +1723,7 @@ let build_tycon loc env tycon_env s subst tycon extenv evdref t = (t,tt) in let b = e_cumul env evdref tt (mkSort s) (* side effect *) in if not b then anomaly (Pp.str "Build_tycon: should be a type"); - { uj_val = EConstr.Unsafe.to_constr t; uj_type = EConstr.Unsafe.to_constr tt } + { uj_val = t; uj_type = tt } (* For a multiple pattern-matching problem Xi on t1..tn with return * type T, [build_inversion_problem Gamma Sigma (t1..tn) T] builds a return @@ -1851,7 +1848,7 @@ let build_inversion_problem loc env sigma tms t = caseloc = loc; casestyle = RegularStyle; typing_function = build_tycon loc env pb_env s subst} in - let pred = EConstr.of_constr (compile pb).uj_val in + let pred = (compile pb).uj_val in (!evdref,pred) (* Here, [pred] is assumed to be in the context built from all *) @@ -1905,7 +1902,7 @@ let extract_arity_signature ?(dolift=true) env0 tomatchl tmsign = let inh_conv_coerce_to_tycon loc env evdref j tycon = match tycon with | Some p -> - let (evd',j) = Coercion.inh_conv_coerce_to true loc env !evdref j (EConstr.of_constr p) in + let (evd',j) = Coercion.inh_conv_coerce_to true loc env !evdref j p in evdref := evd'; j | None -> j @@ -2029,7 +2026,7 @@ let prepare_predicate loc typing_fun env sigma tomatchs arsign tycon pred = let evdref = ref sigma in let predcclj = typing_fun (mk_tycon (EConstr.mkSort newt)) envar evdref rtntyp in let sigma = !evdref in - let predccl = EConstr.of_constr (j_nf_evar sigma predcclj).uj_val in + let predccl = EConstr.of_constr (nf_evar sigma (EConstr.Unsafe.to_constr predcclj.uj_val)) in [sigma, predccl] in List.map @@ -2095,7 +2092,7 @@ let constr_of_pat env evdref arsign pat avoid = let IndType (indf, _) = try find_rectype env ( !evdref) (lift (-(List.length realargs)) ty) with Not_found -> error_case_not_inductive env !evdref - {uj_val = EConstr.Unsafe.to_constr ty; uj_type = Typing.unsafe_type_of env !evdref ty} + {uj_val = ty; uj_type = EConstr.of_constr (Typing.unsafe_type_of env !evdref ty)} in let (ind,u), params = dest_ind_family indf in let params = List.map EConstr.of_constr params in @@ -2297,8 +2294,8 @@ let constrs_of_pats typing_fun env evdref eqns tomatchs sign neqs arity = in let rhs_env = push_rel_context rhs_rels' env in let j = typing_fun (mk_tycon tycon) rhs_env eqn.rhs.it in - let bbody = it_mkLambda_or_LetIn (EConstr.of_constr j.uj_val) rhs_rels' - and btype = it_mkProd_or_LetIn (EConstr.of_constr j.uj_type) rhs_rels' in + let bbody = it_mkLambda_or_LetIn j.uj_val rhs_rels' + and btype = it_mkProd_or_LetIn j.uj_type rhs_rels' in let _btype = evd_comb1 (Typing.type_of env) evdref bbody in let branch_name = Id.of_string ("program_branch_" ^ (string_of_int !i)) in let branch_decl = local_def (Name branch_name, lift !i bbody, lift !i btype) in @@ -2554,10 +2551,10 @@ let compile_program_cases loc style (typing_function, evdref) tycon env let j = compile pb in (* We check for unused patterns *) List.iter (check_unused_pattern env) matx; - let body = it_mkLambda_or_LetIn (applist (EConstr.of_constr j.uj_val, args)) lets in + let body = it_mkLambda_or_LetIn (applist (j.uj_val, args)) lets in let j = - { uj_val = EConstr.Unsafe.to_constr (it_mkLambda_or_LetIn body tomatchs_lets); - uj_type = EConstr.to_constr !evdref tycon; } + { uj_val = it_mkLambda_or_LetIn body tomatchs_lets; + uj_type = EConstr.of_constr (EConstr.to_constr !evdref tycon); } in j (**************************************************************************) @@ -2632,7 +2629,7 @@ let compile_cases loc style (typing_fun, evdref) tycon env (predopt, tomatchl, e let j = compile pb in (* We coerce to the tycon (if an elim predicate was provided) *) - let j = inh_conv_coerce_to_tycon loc env myevdref j (Option.map EConstr.Unsafe.to_constr tycon) in + let j = inh_conv_coerce_to_tycon loc env myevdref j tycon in evdref := !myevdref; j in diff --git a/pretyping/cases.mli b/pretyping/cases.mli index 9016ca5f3..9f26ae9ce 100644 --- a/pretyping/cases.mli +++ b/pretyping/cases.mli @@ -8,9 +8,9 @@ open Names open Term -open EConstr open Evd open Environ +open EConstr open Inductiveops open Glob_term open Evarutil @@ -111,11 +111,11 @@ type 'a pattern_matching_problem = typing_function: type_constraint -> env -> evar_map ref -> 'a option -> unsafe_judgment } -val compile : 'a pattern_matching_problem -> Environ.unsafe_judgment +val compile : 'a pattern_matching_problem -> unsafe_judgment val prepare_predicate : Loc.t -> (Evarutil.type_constraint -> - Environ.env -> Evd.evar_map ref -> 'a -> Environ.unsafe_judgment) -> + Environ.env -> Evd.evar_map ref -> 'a -> unsafe_judgment) -> Environ.env -> Evd.evar_map -> (types * tomatch_type) list -> diff --git a/pretyping/classops.ml b/pretyping/classops.ml index ad43bf322..9011186a3 100644 --- a/pretyping/classops.ml +++ b/pretyping/classops.ml @@ -319,7 +319,7 @@ let coercion_value { coe_value = c; coe_type = t; coe_context = ctx; let subst, ctx = Universes.fresh_universe_context_set_instance ctx in let c' = Vars.subst_univs_level_constr subst c and t' = Vars.subst_univs_level_constr subst t in - (make_judge c' t', b, b'), ctx + (make_judge (EConstr.of_constr c') (EConstr.of_constr t'), b, b'), ctx (* pretty-print functions are now in Pretty *) (* rajouter une coercion dans le graphe *) diff --git a/pretyping/classops.mli b/pretyping/classops.mli index 9fb70534f..a1d030f12 100644 --- a/pretyping/classops.mli +++ b/pretyping/classops.mli @@ -8,9 +8,9 @@ open Names open Term +open Environ open EConstr open Evd -open Environ open Mod_subst (** {6 This is the type of class kinds } *) diff --git a/pretyping/coercion.ml b/pretyping/coercion.ml index b9f14aa43..2d4296fe4 100644 --- a/pretyping/coercion.ml +++ b/pretyping/coercion.ml @@ -53,16 +53,16 @@ let apply_coercion_args env evd check isproj argl funj = let rec apply_rec acc typ = function | [] -> if isproj then - let cst = fst (destConst !evdref (EConstr.of_constr (j_val funj))) in + let cst = fst (destConst !evdref (j_val funj)) in let p = Projection.make cst false in let pb = lookup_projection p env in let args = List.skipn pb.Declarations.proj_npars argl in let hd, tl = match args with hd :: tl -> hd, tl | [] -> assert false in - { uj_val = EConstr.Unsafe.to_constr (applist (mkProj (p, hd), tl)); - uj_type = EConstr.Unsafe.to_constr typ } + { uj_val = applist (mkProj (p, hd), tl); + uj_type = typ } else - { uj_val = EConstr.Unsafe.to_constr (applist (EConstr.of_constr (j_val funj),argl)); - uj_type = EConstr.Unsafe.to_constr typ } + { uj_val = applist (j_val funj,argl); + uj_type = typ } | h::restl -> (* On devrait pouvoir s'arranger pour qu'on n'ait pas a faire hnf_constr *) match EConstr.kind !evdref (EConstr.of_constr (whd_all env !evdref typ)) with | Prod (_,c1,c2) -> @@ -71,7 +71,7 @@ let apply_coercion_args env evd check isproj argl funj = apply_rec (h::acc) (Vars.subst1 h c2) restl | _ -> anomaly (Pp.str "apply_coercion_args") in - let res = apply_rec [] (EConstr.of_constr funj.uj_type) argl in + let res = apply_rec [] funj.uj_type argl in !evdref, res (* appliquer le chemin de coercions de patterns p *) @@ -367,7 +367,7 @@ let apply_coercion env sigma p hj typ_cl = (fun (ja,typ_cl,sigma) i -> let ((fv,isid,isproj),ctx) = coercion_value i in let sigma = Evd.merge_context_set Evd.univ_flexible sigma ctx in - let argl = (class_args_of env sigma typ_cl)@[EConstr.of_constr ja.uj_val] in + let argl = (class_args_of env sigma typ_cl)@[ja.uj_val] in let sigma, jres = apply_coercion_args env sigma true isproj argl fv in @@ -375,7 +375,7 @@ let apply_coercion env sigma p hj typ_cl = { uj_val = ja.uj_val; uj_type = jres.uj_type } else jres), - EConstr.of_constr jres.uj_type,sigma) + jres.uj_type,sigma) (hj,typ_cl,sigma) p in evd, j with NoCoercion as e -> raise e @@ -383,23 +383,23 @@ let apply_coercion env sigma p hj typ_cl = (* Try to coerce to a funclass; raise NoCoercion if not possible *) let inh_app_fun_core env evd j = - let t = whd_all env evd (EConstr.of_constr j.uj_type) in + let t = whd_all env evd j.uj_type in let t = EConstr.of_constr t in match EConstr.kind evd t with | Prod (_,_,_) -> (evd,j) | Evar ev -> let (evd',t) = Evardefine.define_evar_as_product evd ev in - (evd',{ uj_val = j.uj_val; uj_type = EConstr.Unsafe.to_constr t }) + (evd',{ uj_val = j.uj_val; uj_type = t }) | _ -> try let t,p = - lookup_path_to_fun_from env evd (EConstr.of_constr j.uj_type) in + lookup_path_to_fun_from env evd j.uj_type in apply_coercion env evd p j t with Not_found | NoCoercion -> if Flags.is_program_mode () then try let evdref = ref evd in let coercef, t = mu env evdref t in - let res = { uj_val = EConstr.Unsafe.to_constr (app_opt env evdref coercef (EConstr.of_constr j.uj_val)); uj_type = EConstr.Unsafe.to_constr t } in + let res = { uj_val = app_opt env evdref coercef j.uj_val; uj_type = t } in (!evdref, res) with NoSubtacCoercion | NoCoercion -> (evd,j) @@ -415,17 +415,23 @@ let inh_app_fun resolve_tc env evd j = try inh_app_fun_core env (saturate_evd env evd) j with NoCoercion -> (evd, j) +let type_judgment env sigma j = + match EConstr.kind sigma (EConstr.of_constr (whd_all env sigma j.uj_type)) with + | Sort s -> {utj_val = j.uj_val; utj_type = s } + | _ -> error_not_a_type env sigma j + let inh_tosort_force loc env evd j = try - let t,p = lookup_path_to_sort_from env evd (EConstr.of_constr j.uj_type) in + let t,p = lookup_path_to_sort_from env evd j.uj_type in let evd,j1 = apply_coercion env evd p j t in + let whd_evar evd c = EConstr.of_constr (whd_evar evd (EConstr.Unsafe.to_constr c)) in let j2 = on_judgment_type (whd_evar evd) j1 in - (evd,type_judgment env j2) + (evd,type_judgment env evd j2) with Not_found | NoCoercion -> error_not_a_type ~loc env evd j let inh_coerce_to_sort loc env evd j = - let typ = whd_all env evd (EConstr.of_constr j.uj_type) in + let typ = whd_all env evd j.uj_type in match EConstr.kind evd (EConstr.of_constr typ) with | Sort s -> (evd,{ utj_val = j.uj_val; utj_type = s }) | Evar ev -> @@ -437,10 +443,10 @@ let inh_coerce_to_sort loc env evd j = let inh_coerce_to_base loc env evd j = if Flags.is_program_mode () then let evdref = ref evd in - let ct, typ' = mu env evdref (EConstr.of_constr j.uj_type) in + let ct, typ' = mu env evdref j.uj_type in let res = - { uj_val = EConstr.Unsafe.to_constr (app_coercion env evdref ct (EConstr.of_constr j.uj_val)); - uj_type = EConstr.Unsafe.to_constr typ' } + { uj_val = (app_coercion env evdref ct j.uj_val); + uj_type = typ' } in !evdref, res else (evd, j) @@ -463,8 +469,8 @@ let inh_coerce_to_fail env evd rigidonly v t c1 = | Some v -> let evd,j = apply_coercion env evd p - {uj_val = EConstr.Unsafe.to_constr v; uj_type = EConstr.Unsafe.to_constr t} t2 in - evd, Some (EConstr.of_constr j.uj_val), (EConstr.of_constr j.uj_type) + {uj_val = v; uj_type = t} t2 in + evd, Some j.uj_val, j.uj_type | None -> evd, None, t with Not_found -> raise NoCoercion in @@ -510,27 +516,27 @@ let rec inh_conv_coerce_to_fail loc env evd rigidonly v t c1 = let inh_conv_coerce_to_gen resolve_tc rigidonly loc env evd cj t = let (evd', val') = try - inh_conv_coerce_to_fail loc env evd rigidonly (Some (EConstr.of_constr cj.uj_val)) (EConstr.of_constr cj.uj_type) t + inh_conv_coerce_to_fail loc env evd rigidonly (Some cj.uj_val) cj.uj_type t with NoCoercionNoUnifier (best_failed_evd,e) -> try if Flags.is_program_mode () then - coerce_itf loc env evd (Some (EConstr.of_constr cj.uj_val)) (EConstr.of_constr cj.uj_type) t + coerce_itf loc env evd (Some cj.uj_val) cj.uj_type t else raise NoSubtacCoercion with | NoSubtacCoercion when not resolve_tc || not !use_typeclasses_for_conversion -> - error_actual_type ~loc env best_failed_evd cj (EConstr.Unsafe.to_constr t) e + error_actual_type ~loc env best_failed_evd cj t e | NoSubtacCoercion -> let evd' = saturate_evd env evd in try if evd' == evd then - error_actual_type ~loc env best_failed_evd cj (EConstr.Unsafe.to_constr t) e + error_actual_type ~loc env best_failed_evd cj t e else - inh_conv_coerce_to_fail loc env evd' rigidonly (Some (EConstr.of_constr cj.uj_val)) (EConstr.of_constr cj.uj_type) t + inh_conv_coerce_to_fail loc env evd' rigidonly (Some cj.uj_val) cj.uj_type t with NoCoercionNoUnifier (_evd,_error) -> - error_actual_type ~loc env best_failed_evd cj (EConstr.Unsafe.to_constr t) e + error_actual_type ~loc env best_failed_evd cj t e in let val' = match val' with Some v -> v | None -> assert(false) in - (evd',{ uj_val = EConstr.Unsafe.to_constr val'; uj_type = EConstr.Unsafe.to_constr t }) + (evd',{ uj_val = val'; uj_type = t }) let inh_conv_coerce_to resolve_tc = inh_conv_coerce_to_gen resolve_tc false let inh_conv_coerce_rigid_to resolve_tc = inh_conv_coerce_to_gen resolve_tc true diff --git a/pretyping/coercion.mli b/pretyping/coercion.mli index 62d4fb004..bc63d092d 100644 --- a/pretyping/coercion.mli +++ b/pretyping/coercion.mli @@ -10,6 +10,7 @@ open Evd open Names open Term open Environ +open EConstr open Glob_term (** {6 Coercions. } *) @@ -36,7 +37,7 @@ val inh_coerce_to_base : Loc.t -> (** [inh_coerce_to_prod env isevars t] coerces [t] to a product type *) val inh_coerce_to_prod : Loc.t -> - env -> evar_map -> EConstr.types -> evar_map * EConstr.types + env -> evar_map -> types -> evar_map * types (** [inh_conv_coerce_to resolve_tc Loc.t env isevars j t] coerces [j] to an object of type [t]; i.e. it inserts a coercion into [j], if needed, in such @@ -44,16 +45,16 @@ val inh_coerce_to_prod : Loc.t -> applicable. resolve_tc=false disables resolving type classes (as the last resort before failing) *) val inh_conv_coerce_to : bool -> Loc.t -> - env -> evar_map -> unsafe_judgment -> EConstr.types -> evar_map * unsafe_judgment + env -> evar_map -> unsafe_judgment -> types -> evar_map * unsafe_judgment val inh_conv_coerce_rigid_to : bool -> Loc.t -> - env -> evar_map -> unsafe_judgment -> EConstr.types -> evar_map * unsafe_judgment + env -> evar_map -> unsafe_judgment -> types -> evar_map * unsafe_judgment (** [inh_conv_coerces_to loc env isevars t t'] checks if an object of type [t] is coercible to an object of type [t'] adding evar constraints if needed; it fails if no coercion exists *) val inh_conv_coerces_to : Loc.t -> - env -> evar_map -> EConstr.types -> EConstr.types -> evar_map + env -> evar_map -> types -> types -> evar_map (** [inh_pattern_coerce_to loc env isevars pat ind1 ind2] coerces the Cases pattern [pat] typed in [ind1] into a pattern typed in [ind2]; diff --git a/pretyping/evarconv.ml b/pretyping/evarconv.ml index 3b420347b..639d6260e 100644 --- a/pretyping/evarconv.ml +++ b/pretyping/evarconv.ml @@ -1266,8 +1266,8 @@ let solve_unconstrained_impossible_cases env evd = let evd' = Evd.merge_context_set Evd.univ_flexible_alg ~loc evd' ctx in let ty = j_type j in let conv_algo = evar_conv_x full_transparent_state in - let evd' = check_evar_instance evd' evk (EConstr.of_constr ty) conv_algo in - Evd.define evk ty evd' + let evd' = check_evar_instance evd' evk ty conv_algo in + Evd.define evk (EConstr.Unsafe.to_constr ty) evd' | _ -> evd') evd evd let consider_remaining_unif_problems env diff --git a/pretyping/pretype_errors.ml b/pretyping/pretype_errors.ml index 14b25ab36..673554005 100644 --- a/pretyping/pretype_errors.ml +++ b/pretyping/pretype_errors.ml @@ -31,11 +31,13 @@ type position_reporting = (position * int) * EConstr.t type subterm_unification_error = bool * position_reporting * position_reporting * (EConstr.constr * EConstr.constr * unification_error) option +type type_error = (EConstr.constr, EConstr.types) ptype_error + type pretype_error = (* Old Case *) - | CantFindCaseType of constr + | CantFindCaseType of EConstr.constr (* Type inference unification *) - | ActualTypeNotCoercible of unsafe_judgment * types * unification_error + | ActualTypeNotCoercible of EConstr.unsafe_judgment * EConstr.types * unification_error (* Tactic unification *) | UnifOccurCheck of existential_key * EConstr.constr | UnsolvableImplicit of existential_key * Evd.unsolvability_explanation option @@ -50,7 +52,7 @@ type pretype_error = | NonLinearUnification of Name.t * EConstr.constr (* Pretyping *) | VarNotFound of Id.t - | UnexpectedType of constr * constr + | UnexpectedType of EConstr.constr * EConstr.constr | NotProduct of EConstr.constr | TypingError of type_error | CannotUnifyOccurrences of subterm_unification_error @@ -75,14 +77,19 @@ let error_actual_type ?loc env sigma {uj_val=c;uj_type=actty} expty reason = raise_pretype_error ?loc (env, sigma, ActualTypeNotCoercible (j, expty, reason)) +let error_actual_type_core ?loc env sigma {uj_val=c;uj_type=actty} expty = + let j = {uj_val=c;uj_type=actty} in + raise_type_error ?loc + (env, sigma, ActualType (j, expty)) + let error_cant_apply_not_functional ?loc env sigma rator randl = raise_type_error ?loc - (env, sigma, CantApplyNonFunctional (rator, Array.of_list randl)) + (env, sigma, CantApplyNonFunctional (rator, randl)) let error_cant_apply_bad_type ?loc env sigma (n,c,t) rator randl = raise_type_error ?loc (env, sigma, - CantApplyBadType ((n,c,t), rator, Array.of_list randl)) + CantApplyBadType ((n,c,t), rator, randl)) let error_ill_formed_branch ?loc env sigma c i actty expty = raise_type_error @@ -98,9 +105,16 @@ let error_ill_typed_rec_body ?loc env sigma i na jl tys = raise_type_error ?loc (env, sigma, IllTypedRecBody (i, na, jl, tys)) +let error_elim_arity ?loc env sigma pi s c j a = + raise_type_error ?loc + (env, sigma, ElimArity (pi, s, c, j, a)) + let error_not_a_type ?loc env sigma j = raise_type_error ?loc (env, sigma, NotAType j) +let error_assumption ?loc env sigma j = + raise_type_error ?loc (env, sigma, BadAssumption j) + (*s Implicit arguments synthesis errors. It is hard to find a precise location. *) diff --git a/pretyping/pretype_errors.mli b/pretyping/pretype_errors.mli index 2e707a0ff..0ebe4817c 100644 --- a/pretyping/pretype_errors.mli +++ b/pretyping/pretype_errors.mli @@ -32,11 +32,13 @@ type position_reporting = (position * int) * EConstr.t type subterm_unification_error = bool * position_reporting * position_reporting * (EConstr.constr * EConstr.constr * unification_error) option +type type_error = (EConstr.constr, EConstr.types) ptype_error + type pretype_error = (** Old Case *) - | CantFindCaseType of constr + | CantFindCaseType of EConstr.constr (** Type inference unification *) - | ActualTypeNotCoercible of unsafe_judgment * types * unification_error + | ActualTypeNotCoercible of EConstr.unsafe_judgment * EConstr.types * unification_error (** Tactic Unification *) | UnifOccurCheck of existential_key * EConstr.constr | UnsolvableImplicit of existential_key * Evd.unsolvability_explanation option @@ -51,7 +53,7 @@ type pretype_error = | NonLinearUnification of Name.t * EConstr.constr (** Pretyping *) | VarNotFound of Id.t - | UnexpectedType of constr * constr + | UnexpectedType of EConstr.constr * EConstr.constr | NotProduct of EConstr.constr | TypingError of type_error | CannotUnifyOccurrences of subterm_unification_error @@ -65,34 +67,45 @@ val precatchable_exception : exn -> bool (** Raising errors *) val error_actual_type : - ?loc:Loc.t -> env -> Evd.evar_map -> unsafe_judgment -> constr -> + ?loc:Loc.t -> env -> Evd.evar_map -> EConstr.unsafe_judgment -> EConstr.constr -> unification_error -> 'b +val error_actual_type_core : + ?loc:Loc.t -> env -> Evd.evar_map -> EConstr.unsafe_judgment -> EConstr.constr -> 'b + val error_cant_apply_not_functional : ?loc:Loc.t -> env -> Evd.evar_map -> - unsafe_judgment -> unsafe_judgment list -> 'b + EConstr.unsafe_judgment -> EConstr.unsafe_judgment array -> 'b val error_cant_apply_bad_type : - ?loc:Loc.t -> env -> Evd.evar_map -> int * constr * constr -> - unsafe_judgment -> unsafe_judgment list -> 'b + ?loc:Loc.t -> env -> Evd.evar_map -> int * EConstr.constr * EConstr.constr -> + EConstr.unsafe_judgment -> EConstr.unsafe_judgment array -> 'b val error_case_not_inductive : - ?loc:Loc.t -> env -> Evd.evar_map -> unsafe_judgment -> 'b + ?loc:Loc.t -> env -> Evd.evar_map -> EConstr.unsafe_judgment -> 'b val error_ill_formed_branch : ?loc:Loc.t -> env -> Evd.evar_map -> - constr -> pconstructor -> constr -> constr -> 'b + EConstr.constr -> pconstructor -> EConstr.constr -> EConstr.constr -> 'b val error_number_branches : ?loc:Loc.t -> env -> Evd.evar_map -> - unsafe_judgment -> int -> 'b + EConstr.unsafe_judgment -> int -> 'b val error_ill_typed_rec_body : ?loc:Loc.t -> env -> Evd.evar_map -> - int -> Name.t array -> unsafe_judgment array -> types array -> 'b + int -> Name.t array -> EConstr.unsafe_judgment array -> EConstr.types array -> 'b + +val error_elim_arity : + ?loc:Loc.t -> env -> Evd.evar_map -> + pinductive -> sorts_family list -> EConstr.constr -> + EConstr.unsafe_judgment -> (sorts_family * sorts_family * arity_error) option -> 'b val error_not_a_type : - ?loc:Loc.t -> env -> Evd.evar_map -> unsafe_judgment -> 'b + ?loc:Loc.t -> env -> Evd.evar_map -> EConstr.unsafe_judgment -> 'b + +val error_assumption : + ?loc:Loc.t -> env -> Evd.evar_map -> EConstr.unsafe_judgment -> 'b val error_cannot_coerce : env -> Evd.evar_map -> EConstr.constr * EConstr.constr -> 'b @@ -124,12 +137,12 @@ val error_non_linear_unification : env -> Evd.evar_map -> (** {6 Ml Case errors } *) val error_cant_find_case_type : - ?loc:Loc.t -> env -> Evd.evar_map -> constr -> 'b + ?loc:Loc.t -> env -> Evd.evar_map -> EConstr.constr -> 'b (** {6 Pretyping errors } *) val error_unexpected_type : - ?loc:Loc.t -> env -> Evd.evar_map -> constr -> constr -> 'b + ?loc:Loc.t -> env -> Evd.evar_map -> EConstr.constr -> EConstr.constr -> 'b val error_not_product : ?loc:Loc.t -> env -> Evd.evar_map -> EConstr.constr -> 'b diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml index cac31a1c5..49a0bccee 100644 --- a/pretyping/pretyping.ml +++ b/pretyping/pretyping.ml @@ -34,6 +34,7 @@ open Reductionops open Environ open Type_errors open Typeops +open Typing open Globnames open Nameops open Evarutil @@ -124,7 +125,7 @@ let e_new_evar env evdref ?src ?naming typ = let sigma = Sigma.Unsafe.of_evar_map !evdref in let Sigma (e, sigma, _) = new_evar_instance sign sigma typ' ?src ?naming instance in evdref := Sigma.to_evar_map sigma; - e + EConstr.of_constr e let push_rec_types (lna,typarray,_) env = let ctxt = Array.map2_i (fun i na t -> local_assum (na, lift i t)) lna typarray in @@ -425,21 +426,19 @@ let invert_ltac_bound_name lvar env id0 id = str " which is not bound in current context.") let protected_get_type_of env sigma c = - try Retyping.get_type_of ~lax:true env.ExtraEnv.env sigma c + try EConstr.of_constr (Retyping.get_type_of ~lax:true env.ExtraEnv.env sigma c) with Retyping.RetypeError _ -> user_err (str "Cannot reinterpret " ++ quote (print_constr (EConstr.Unsafe.to_constr c)) ++ str " in the current environment.") -let j_val j = EConstr.of_constr (j_val j) - let pretype_id pretype k0 loc env evdref lvar id = let sigma = !evdref in (* Look for the binder of [id] *) try let (n,_,typ) = lookup_rel_id id (rel_context env) in let typ = EConstr.of_constr typ in - { uj_val = EConstr.Unsafe.to_constr (mkRel n); uj_type = EConstr.Unsafe.to_constr (lift n typ) } + { uj_val = mkRel n; uj_type = lift n typ } with Not_found -> let env = ltac_interp_name_env k0 lvar env in (* Check if [id] is an ltac variable *) @@ -447,7 +446,7 @@ let pretype_id pretype k0 loc env evdref lvar id = let (ids,c) = Id.Map.find id lvar.ltac_constrs in let subst = List.map (invert_ltac_bound_name lvar env id) ids in let c = substl subst (EConstr.of_constr c) in - { uj_val = EConstr.Unsafe.to_constr c; uj_type = protected_get_type_of env sigma c } + { uj_val = c; uj_type = protected_get_type_of env sigma c } with Not_found -> try let {closure;term} = Id.Map.find id lvar.ltac_uconstrs in let lvar = { @@ -472,7 +471,7 @@ let pretype_id pretype k0 loc env evdref lvar id = end; (* Check if [id] is a section or goal variable *) try - { uj_val = Constr.mkVar id; uj_type = NamedDecl.get_type (lookup_named id env) } + { uj_val = mkVar id; uj_type = EConstr.of_constr (NamedDecl.get_type (lookup_named id env)) } with Not_found -> (* [id] not found, standard error message *) error_var_not_found ~loc id @@ -511,9 +510,6 @@ let pretype_global loc rigid env evd gr us = let (sigma, c) = Evd.fresh_global ~loc ~rigid ?names:instance env.ExtraEnv.env evd gr in (sigma, EConstr.of_constr c) -let make_judge c t = - make_judge (EConstr.Unsafe.to_constr c) (EConstr.Unsafe.to_constr t) - let pretype_ref loc evdref env ref us = match ref with | VarRef id -> @@ -527,14 +523,14 @@ let pretype_ref loc evdref env ref us = | ref -> let evd, c = pretype_global loc univ_flexible env !evdref ref us in let () = evdref := evd in - let ty = Typing.unsafe_type_of env.ExtraEnv.env evd c in + let ty = unsafe_type_of env.ExtraEnv.env evd c in let ty = EConstr.of_constr ty in make_judge c ty let judge_of_Type loc evd s = let evd, s = interp_universe ~loc evd s in let judge = - { uj_val = Constr.mkSort (Type s); uj_type = Constr.mkSort (Type (Univ.super s)) } + { uj_val = mkSort (Type s); uj_type = mkSort (Type (Univ.super s)) } in evd, judge @@ -550,7 +546,7 @@ let new_type_evar env evdref loc = univ_flexible_alg ~src:(loc,Evar_kinds.InternalHole) in evdref := Sigma.to_evar_map sigma; - e + EConstr.of_constr e let (f_genarg_interp, genarg_interp_hook) = Hook.make () @@ -591,25 +587,25 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre let env = ltac_interp_name_env k0 lvar env in let ty = match tycon with - | Some ty -> EConstr.Unsafe.to_constr ty + | Some ty -> ty | None -> new_type_evar env evdref loc in let k = Evar_kinds.MatchingVar (someta,n) in - { uj_val = e_new_evar env evdref ~src:(loc,k) (EConstr.of_constr ty); uj_type = ty } + { uj_val = e_new_evar env evdref ~src:(loc,k) ty; uj_type = ty } | GHole (loc, k, naming, None) -> let env = ltac_interp_name_env k0 lvar env in let ty = match tycon with - | Some ty -> EConstr.Unsafe.to_constr ty + | Some ty -> ty | None -> new_type_evar env evdref loc in - { uj_val = e_new_evar env evdref ~src:(loc,k) ~naming (EConstr.of_constr ty); uj_type = ty } + { uj_val = e_new_evar env evdref ~src:(loc,k) ~naming ty; uj_type = ty } | GHole (loc, k, _naming, Some arg) -> let env = ltac_interp_name_env k0 lvar env in let ty = match tycon with - | Some ty -> EConstr.Unsafe.to_constr ty + | Some ty -> ty | None -> new_type_evar env evdref loc in let ist = lvar.ltac_genargs in @@ -622,14 +618,14 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre [] -> ctxt | (na,bk,None,ty)::bl -> let ty' = pretype_type empty_valcon env evdref lvar ty in - let dcl = LocalAssum (na, ty'.utj_val) in - let dcl' = LocalAssum (ltac_interp_name lvar na,ty'.utj_val) in + let dcl = local_assum (na, ty'.utj_val) in + let dcl' = local_assum (ltac_interp_name lvar na,ty'.utj_val) in type_bl (push_rel dcl env) (Context.Rel.add dcl' ctxt) bl | (na,bk,Some bd,ty)::bl -> let ty' = pretype_type empty_valcon env evdref lvar ty in - let bd' = pretype (mk_tycon (EConstr.of_constr ty'.utj_val)) env evdref lvar bd in - let dcl = LocalDef (na, bd'.uj_val, ty'.utj_val) in - let dcl' = LocalDef (ltac_interp_name lvar na, bd'.uj_val, ty'.utj_val) in + let bd' = pretype (mk_tycon ty'.utj_val) env evdref lvar bd in + let dcl = local_def (na, bd'.uj_val, ty'.utj_val) in + let dcl' = local_def (ltac_interp_name lvar na, bd'.uj_val, ty'.utj_val) in type_bl (push_rel dcl env) (Context.Rel.add dcl' ctxt) bl in let ctxtv = Array.map (type_bl env Context.Rel.empty) bl in let larj = @@ -637,8 +633,8 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre (fun e ar -> pretype_type empty_valcon (push_rel_context e env) evdref lvar ar) ctxtv lar in - let lara = Array.map (fun a -> EConstr.of_constr a.utj_val) larj in - let ftys = Array.map2 (fun e a -> EConstr.it_mkProd_or_LetIn a e) ctxtv lara in + let lara = Array.map (fun a -> a.utj_val) larj in + let ftys = Array.map2 (fun e a -> it_mkProd_or_LetIn a e) ctxtv lara in let nbfix = Array.length lar in let names = Array.map (fun id -> Name id) names in let _ = @@ -662,8 +658,8 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre (lift nbfix ftys.(i)) in let nenv = push_rel_context ctxt newenv in let j = pretype (mk_tycon ty) nenv evdref lvar def in - { uj_val = Termops.it_mkLambda_or_LetIn j.uj_val ctxt; - uj_type = Termops.it_mkProd_or_LetIn j.uj_type ctxt }) + { uj_val = it_mkLambda_or_LetIn j.uj_val ctxt; + uj_type = it_mkProd_or_LetIn j.uj_type ctxt }) ctxtv vdef in Typing.check_type_fixpoint loc env.ExtraEnv.env evdref names ftys vdefj; let nf c = nf_evar !evdref c in @@ -714,11 +710,11 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre (* Bidirectional typechecking hint: parameters of a constructor are completely determined by a typing constraint *) - if Flags.is_program_mode () && length > 0 && isConstruct !evdref (EConstr.of_constr fj.uj_val) then + if Flags.is_program_mode () && length > 0 && isConstruct !evdref fj.uj_val then match tycon with | None -> [] | Some ty -> - let ((ind, i), u) = destConstruct !evdref (EConstr.of_constr fj.uj_val) in + let ((ind, i), u) = destConstruct !evdref fj.uj_val in let npars = inductive_nparams ind in if Int.equal npars 0 then [] else @@ -731,7 +727,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre else [] in let app_f = - match EConstr.kind !evdref (EConstr.of_constr fj.uj_val) with + match EConstr.kind !evdref fj.uj_val with | Const (p, u) when Environ.is_projection p env.ExtraEnv.env -> let p = Projection.make p false in let pb = Environ.lookup_projection p env.ExtraEnv.env in @@ -746,7 +742,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre | c::rest -> let argloc = loc_of_glob_constr c in let resj = evd_comb1 (Coercion.inh_app_fun resolve_tc env.ExtraEnv.env) evdref resj in - let resty = whd_all env.ExtraEnv.env !evdref (EConstr.of_constr resj.uj_type) in + let resty = whd_all env.ExtraEnv.env !evdref resj.uj_type in let resty = EConstr.of_constr resty in match EConstr.kind !evdref resty with | Prod (na,c1,c2) -> @@ -761,18 +757,18 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre else [], j_val hj in let value, typ = app_f n (j_val resj) ujval, subst1 ujval c2 in - let j = { uj_val = EConstr.Unsafe.to_constr value; uj_type = EConstr.Unsafe.to_constr typ } in + let j = { uj_val = value; uj_type = typ } in apply_rec env (n+1) j candargs rest | _ -> let hj = pretype empty_tycon env evdref lvar c in error_cant_apply_not_functional ~loc:(Loc.merge floc argloc) env.ExtraEnv.env !evdref - resj [hj] + resj [|hj|] in let resj = apply_rec env 1 fj candargs args in let resj = - match EConstr.kind !evdref (EConstr.of_constr resj.uj_val) with + match EConstr.kind !evdref resj.uj_val with | App (f,args) -> if is_template_polymorphic env.ExtraEnv.env !evdref f then (* Special case for inductive type applications that must be @@ -804,7 +800,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre (* The name specified by ltac is used also to create bindings. So the substitution must also be applied on variables before they are looked up in the rel context. *) - let var = LocalAssum (name, j.utj_val) in + let var = local_assum (name, j.utj_val) in let j' = pretype rng (push_rel var env) evdref lvar c2 in let name = ltac_interp_name lvar name in let resj = judge_of_abstraction env.ExtraEnv.env (orelse_name name name') j j' in @@ -818,9 +814,9 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre let j' = match name with | Anonymous -> let j = pretype_type empty_valcon env evdref lvar c2 in - { j with utj_val = EConstr.Unsafe.to_constr (lift 1 (EConstr.of_constr j.utj_val)) } + { j with utj_val = lift 1 j.utj_val } | Name _ -> - let var = LocalAssum (name, j.utj_val) in + let var = local_assum (name, j.utj_val) in let env' = push_rel var env in pretype_type empty_valcon env' evdref lvar c2 in @@ -839,27 +835,27 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre match c1 with | GCast (loc, c, CastConv t) -> let tj = pretype_type empty_valcon env evdref lvar t in - pretype (mk_tycon (EConstr.of_constr tj.utj_val)) env evdref lvar c + pretype (mk_tycon tj.utj_val) env evdref lvar c | _ -> pretype empty_tycon env evdref lvar c1 in let t = evd_comb1 (Evarsolve.refresh_universes ~onlyalg:true ~status:Evd.univ_flexible (Some false) env.ExtraEnv.env) - evdref (EConstr.of_constr j.uj_type) in + evdref j.uj_type in + let t = EConstr.of_constr t in (* The name specified by ltac is used also to create bindings. So the substitution must also be applied on variables before they are looked up in the rel context. *) - let var = LocalDef (name, j.uj_val, t) in - let t = EConstr.of_constr t in + let var = local_def (name, j.uj_val, t) in let tycon = lift_tycon 1 tycon in let j' = pretype tycon (push_rel var env) evdref lvar c2 in let name = ltac_interp_name lvar name in - { uj_val = EConstr.Unsafe.to_constr (mkLetIn (name, EConstr.of_constr j.uj_val, t, EConstr.of_constr j'.uj_val)) ; - uj_type = EConstr.Unsafe.to_constr (subst1 (EConstr.of_constr j.uj_val) (EConstr.of_constr j'.uj_type)) } + { uj_val = mkLetIn (name, j.uj_val, t, j'.uj_val) ; + uj_type = subst1 j.uj_val j'.uj_type } | GLetTuple (loc,nal,(na,po),c,d) -> let cj = pretype empty_tycon env evdref lvar c in let (IndType (indf,realargs)) = - try find_rectype env.ExtraEnv.env !evdref (EConstr.of_constr cj.uj_type) + try find_rectype env.ExtraEnv.env !evdref cj.uj_type with Not_found -> let cloc = loc_of_glob_constr c in error_case_not_inductive ~loc:cloc env.ExtraEnv.env !evdref cj @@ -883,7 +879,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre | na :: names, (LocalAssum (_,t) :: l) -> let t = EConstr.of_constr t in let proj = Projection.make ps.(cs.cs_nargs - k) true in - local_def (na, lift (cs.cs_nargs - n) (mkProj (proj, EConstr.of_constr cj.uj_val)), t) + local_def (na, lift (cs.cs_nargs - n) (mkProj (proj, cj.uj_val)), t) :: aux (n+1) (k + 1) names l | na :: names, (decl :: l) -> set_name na decl :: aux (n+1) k names l @@ -897,7 +893,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre let fsign = List.map2 set_name nal fsign in let f = it_mkLambda_or_LetIn f fsign in let ci = make_case_info env.ExtraEnv.env (fst ind) LetStyle in - mkCase (ci, p, EConstr.of_constr cj.uj_val,[|f|]) + mkCase (ci, p, cj.uj_val,[|f|]) else it_mkLambda_or_LetIn f fsign in let env_f = push_rel_context fsign env in @@ -914,7 +910,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre | Some p -> let env_p = push_rel_context psign env in let pj = pretype_type empty_valcon env_p evdref lvar p in - let ccl = nf_evar !evdref (EConstr.of_constr pj.utj_val) in + let ccl = nf_evar !evdref pj.utj_val in let psign = make_arity_signature env.ExtraEnv.env true indf in (* with names *) let p = it_mkLambda_or_LetIn ccl psign in let inst = @@ -922,18 +918,19 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre @[EConstr.of_constr (build_dependent_constructor cs)] in let lp = lift cs.cs_nargs p in let fty = hnf_lam_applist env.ExtraEnv.env !evdref lp inst in - let fj = pretype (mk_tycon (EConstr.of_constr fty)) env_f evdref lvar d in + let fty = EConstr.of_constr fty in + let fj = pretype (mk_tycon fty) env_f evdref lvar d in let v = let ind,_ = dest_ind_family indf in - Typing.check_allowed_sort env.ExtraEnv.env !evdref ind (EConstr.of_constr cj.uj_val) p; - obj ind p cj.uj_val (EConstr.of_constr fj.uj_val) + Typing.check_allowed_sort env.ExtraEnv.env !evdref ind cj.uj_val p; + obj ind p cj.uj_val fj.uj_val in - { uj_val = EConstr.Unsafe.to_constr v; uj_type = EConstr.Unsafe.to_constr (substl (realargs@[EConstr.of_constr cj.uj_val]) ccl) } + { uj_val = v; uj_type = (substl (realargs@[cj.uj_val]) ccl) } | None -> let tycon = lift_tycon cs.cs_nargs tycon in let fj = pretype tycon env_f evdref lvar d in - let ccl = nf_evar !evdref (EConstr.of_constr fj.uj_type) in + let ccl = nf_evar !evdref fj.uj_type in let ccl = if noccur_between !evdref 1 cs.cs_nargs ccl then lift (- cs.cs_nargs) ccl @@ -944,14 +941,14 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre let p = it_mkLambda_or_LetIn (lift (nar+1) ccl) psign in let v = let ind,_ = dest_ind_family indf in - Typing.check_allowed_sort env.ExtraEnv.env !evdref ind (EConstr.of_constr cj.uj_val) p; - obj ind p cj.uj_val (EConstr.of_constr fj.uj_val) - in { uj_val = EConstr.Unsafe.to_constr v; uj_type = EConstr.Unsafe.to_constr ccl }) + Typing.check_allowed_sort env.ExtraEnv.env !evdref ind cj.uj_val p; + obj ind p cj.uj_val fj.uj_val + in { uj_val = v; uj_type = ccl }) | GIf (loc,c,(na,po),b1,b2) -> let cj = pretype empty_tycon env evdref lvar c in let (IndType (indf,realargs)) = - try find_rectype env.ExtraEnv.env !evdref (EConstr.of_constr cj.uj_type) + try find_rectype env.ExtraEnv.env !evdref cj.uj_type with Not_found -> let cloc = loc_of_glob_constr c in error_case_not_inductive ~loc:cloc env.ExtraEnv.env !evdref cj in @@ -973,16 +970,16 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre | Some p -> let env_p = push_rel_context psign env in let pj = pretype_type empty_valcon env_p evdref lvar p in - let ccl = nf_evar !evdref (EConstr.of_constr pj.utj_val) in + let ccl = nf_evar !evdref pj.utj_val in let pred = it_mkLambda_or_LetIn ccl psign in - let typ = lift (- nar) (EConstr.of_constr (beta_applist !evdref (pred,[EConstr.of_constr cj.uj_val]))) in + let typ = lift (- nar) (EConstr.of_constr (beta_applist !evdref (pred,[cj.uj_val]))) in pred, typ | None -> let p = match tycon with | Some ty -> ty | None -> let env = ltac_interp_name_env k0 lvar env in - EConstr.of_constr (new_type_evar env evdref loc) + new_type_evar env evdref loc in it_mkLambda_or_LetIn (lift (nar+1) p) psign, p in let pred = nf_evar !evdref pred in @@ -991,6 +988,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre let n = Context.Rel.length cs.cs_args in let pi = lift n pred in (* liftn n 2 pred ? *) let pi = beta_applist !evdref (pi, [EConstr.of_constr (build_dependent_constructor cs)]) in + let pi = EConstr.of_constr pi in let csgn = if not !allow_anonymous_refs then List.map (set_name Anonymous) cs.cs_args @@ -1000,18 +998,18 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre cs.cs_args in let env_c = push_rel_context csgn env in - let bj = pretype (mk_tycon (EConstr.of_constr pi)) env_c evdref lvar b in - it_mkLambda_or_LetIn (EConstr.of_constr bj.uj_val) cs.cs_args in + let bj = pretype (mk_tycon pi) env_c evdref lvar b in + it_mkLambda_or_LetIn bj.uj_val cs.cs_args in let b1 = f cstrs.(0) b1 in let b2 = f cstrs.(1) b2 in let v = let ind,_ = dest_ind_family indf in let ci = make_case_info env.ExtraEnv.env (fst ind) IfStyle in let pred = nf_evar !evdref pred in - Typing.check_allowed_sort env.ExtraEnv.env !evdref ind (EConstr.of_constr cj.uj_val) pred; - mkCase (ci, pred, EConstr.of_constr cj.uj_val, [|b1;b2|]) + Typing.check_allowed_sort env.ExtraEnv.env !evdref ind cj.uj_val pred; + mkCase (ci, pred, cj.uj_val, [|b1;b2|]) in - let cj = { uj_val = EConstr.Unsafe.to_constr v; uj_type = EConstr.Unsafe.to_constr p } in + let cj = { uj_val = v; uj_type = p } in inh_conv_coerce_to_tycon loc env evdref cj tycon | GCases (loc,sty,po,tml,eqns) -> @@ -1030,36 +1028,36 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre let tj = pretype_type empty_valcon env evdref lvar t in let tval = evd_comb1 (Evarsolve.refresh_universes ~onlyalg:true ~status:Evd.univ_flexible (Some false) env.ExtraEnv.env) - evdref (EConstr.of_constr tj.utj_val) in + evdref tj.utj_val in let tval = EConstr.of_constr tval in let tval = nf_evar !evdref tval in let cj, tval = match k with | VMcast -> let cj = pretype empty_tycon env evdref lvar c in - let cty = nf_evar !evdref (EConstr.of_constr cj.uj_type) and tval = nf_evar !evdref tval in + let cty = nf_evar !evdref cj.uj_type and tval = nf_evar !evdref tval in if not (occur_existential !evdref cty || occur_existential !evdref tval) then let (evd,b) = Reductionops.vm_infer_conv env.ExtraEnv.env !evdref cty tval in if b then (evdref := evd; cj, tval) else - error_actual_type ~loc env.ExtraEnv.env !evdref cj (EConstr.Unsafe.to_constr tval) + error_actual_type ~loc env.ExtraEnv.env !evdref cj tval (ConversionFailed (env.ExtraEnv.env,cty,tval)) else user_err ~loc (str "Cannot check cast with vm: " ++ str "unresolved arguments remain.") | NATIVEcast -> let cj = pretype empty_tycon env evdref lvar c in - let cty = nf_evar !evdref (EConstr.of_constr cj.uj_type) and tval = nf_evar !evdref tval in + let cty = nf_evar !evdref cj.uj_type and tval = nf_evar !evdref tval in begin let (evd,b) = Nativenorm.native_infer_conv env.ExtraEnv.env !evdref cty tval in if b then (evdref := evd; cj, tval) else - error_actual_type ~loc env.ExtraEnv.env !evdref cj (EConstr.Unsafe.to_constr tval) + error_actual_type ~loc env.ExtraEnv.env !evdref cj tval (ConversionFailed (env.ExtraEnv.env,cty,tval)) end | _ -> pretype (mk_tycon tval) env evdref lvar c, tval in - let v = mkCast (EConstr.of_constr cj.uj_val, k, tval) in - { uj_val = EConstr.Unsafe.to_constr v; uj_type = EConstr.Unsafe.to_constr tval } + let v = mkCast (cj.uj_val, k, tval) in + { uj_val = v; uj_type = tval } in inh_conv_coerce_to_tycon loc env evdref cj tycon and pretype_instance k0 resolve_tc env evdref lvar loc hyps evk update = @@ -1070,7 +1068,7 @@ and pretype_instance k0 resolve_tc env evdref lvar loc hyps evk update = try let c = List.assoc id update in let c = pretype k0 resolve_tc (mk_tycon t) env evdref lvar c in - EConstr.of_constr c.uj_val, List.remove_assoc id update + c.uj_val, List.remove_assoc id update with Not_found -> try let (n,_,t') = lookup_rel_id id (rel_context env) in @@ -1103,13 +1101,14 @@ and pretype_type k0 resolve_tc valcon (env : ExtraEnv.t) evdref lvar = function let s = let sigma = !evdref in let t = Retyping.get_type_of env.ExtraEnv.env sigma v in - match EConstr.kind sigma (EConstr.of_constr (whd_all env.ExtraEnv.env sigma (EConstr.of_constr t))) with + let t = EConstr.of_constr t in + match EConstr.kind sigma (EConstr.of_constr (whd_all env.ExtraEnv.env sigma t)) with | Sort s -> s | Evar ev when is_Type (existential_type sigma ev) -> evd_comb1 (define_evar_as_sort env.ExtraEnv.env) evdref ev | _ -> anomaly (Pp.str "Found a type constraint which is not a type") in - { utj_val = EConstr.Unsafe.to_constr v; + { utj_val = v; utj_type = s } | None -> let env = ltac_interp_name_env k0 lvar env in @@ -1123,10 +1122,10 @@ and pretype_type k0 resolve_tc valcon (env : ExtraEnv.t) evdref lvar = function match valcon with | None -> tj | Some v -> - if e_cumul env.ExtraEnv.env evdref v (EConstr.of_constr tj.utj_val) then tj + if e_cumul env.ExtraEnv.env evdref v tj.utj_val then tj else error_unexpected_type - ~loc:(loc_of_glob_constr c) env.ExtraEnv.env !evdref tj.utj_val (EConstr.Unsafe.to_constr v) + ~loc:(loc_of_glob_constr c) env.ExtraEnv.env !evdref tj.utj_val v let ise_pretype_gen flags env sigma lvar kind c = let env = make_env env in @@ -1140,7 +1139,7 @@ let ise_pretype_gen flags env sigma lvar kind c = | IsType -> (pretype_type k0 flags.use_typeclasses empty_valcon env evdref lvar c).utj_val in - process_inference_flags flags env.ExtraEnv.env sigma (!evdref,EConstr.of_constr c') + process_inference_flags flags env.ExtraEnv.env sigma (!evdref,c') let default_inference_flags fail = { use_typeclasses = true; @@ -1167,9 +1166,9 @@ let empty_lvar : ltac_var_map = { } let on_judgment sigma f j = - let c = mkCast(EConstr.of_constr j.uj_val,DEFAULTcast, EConstr.of_constr j.uj_type) in + let c = mkCast(j.uj_val,DEFAULTcast, j.uj_type) in let (c,_,t) = destCast sigma (f c) in - {uj_val = EConstr.Unsafe.to_constr c; uj_type = EConstr.Unsafe.to_constr t} + {uj_val = c; uj_type = t} let understand_judgment env sigma c = let env = make_env env in diff --git a/pretyping/pretyping.mli b/pretyping/pretyping.mli index 603b9f9ea..2f3ce3afa 100644 --- a/pretyping/pretyping.mli +++ b/pretyping/pretyping.mli @@ -110,11 +110,11 @@ val understand : ?flags:inference_flags -> ?expected_type:typing_constraint -> (** Idem but returns the judgment of the understood term *) val understand_judgment : env -> evar_map -> - glob_constr -> unsafe_judgment Evd.in_evar_universe_context + glob_constr -> EConstr.unsafe_judgment Evd.in_evar_universe_context (** Idem but do not fail on unresolved evars (type cl*) val understand_judgment_tcc : env -> evar_map ref -> - glob_constr -> unsafe_judgment + glob_constr -> EConstr.unsafe_judgment val type_uconstr : ?flags:inference_flags -> @@ -145,11 +145,11 @@ val check_evars : env -> evar_map -> evar_map -> EConstr.constr -> unit (** Internal of Pretyping... *) val pretype : int -> bool -> type_constraint -> env -> evar_map ref -> - ltac_var_map -> glob_constr -> unsafe_judgment + ltac_var_map -> glob_constr -> EConstr.unsafe_judgment val pretype_type : int -> bool -> val_constraint -> env -> evar_map ref -> - ltac_var_map -> glob_constr -> unsafe_type_judgment + ltac_var_map -> glob_constr -> EConstr.unsafe_type_judgment val ise_pretype_gen : inference_flags -> env -> evar_map -> @@ -163,5 +163,5 @@ val interp_sort : ?loc:Loc.t -> evar_map -> glob_sort -> evar_map * sorts val interp_elimination_sort : glob_sort -> sorts_family val genarg_interp_hook : - (types -> env -> evar_map -> unbound_ltac_var_map -> - Genarg.glob_generic_argument -> constr * evar_map) Hook.t + (EConstr.types -> env -> evar_map -> unbound_ltac_var_map -> + Genarg.glob_generic_argument -> EConstr.constr * evar_map) Hook.t diff --git a/pretyping/retyping.ml b/pretyping/retyping.ml index 2efb02417..a7ccf98a6 100644 --- a/pretyping/retyping.ml +++ b/pretyping/retyping.ml @@ -244,7 +244,7 @@ let get_type_of ?(polyprop=true) ?(lax=false) env sigma c = if lax then EConstr.Unsafe.to_constr (f env c) else EConstr.Unsafe.to_constr (anomaly_on_error (f env) c) (* Makes an unsafe judgment from a constr *) -let get_judgment_of env evc c = { uj_val = EConstr.Unsafe.to_constr c; uj_type = get_type_of env evc c } +let get_judgment_of env evc c = { uj_val = c; uj_type = EConstr.of_constr (get_type_of env evc c) } (* Returns sorts of a context *) let sorts_of_context env evc ctxt = diff --git a/pretyping/retyping.mli b/pretyping/retyping.mli index 08f750287..c84403890 100644 --- a/pretyping/retyping.mli +++ b/pretyping/retyping.mli @@ -35,7 +35,7 @@ val get_sort_family_of : ?polyprop:bool -> env -> evar_map -> EConstr.types -> sorts_family (** Makes an unsafe judgment from a constr *) -val get_judgment_of : env -> evar_map -> EConstr.constr -> unsafe_judgment +val get_judgment_of : env -> evar_map -> EConstr.constr -> EConstr.unsafe_judgment val type_of_global_reference_knowing_parameters : env -> evar_map -> EConstr.constr -> EConstr.constr array -> types diff --git a/pretyping/typing.ml b/pretyping/typing.ml index c948f9b9a..17adea5f2 100644 --- a/pretyping/typing.ml +++ b/pretyping/typing.ml @@ -6,20 +6,31 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +module CVars = Vars + open Pp open CErrors open Util open Term +open EConstr open Vars open Environ open Reductionops -open Type_errors open Inductive open Inductiveops open Typeops open Arguments_renaming +open Pretype_errors open Context.Rel.Declaration +let local_assum (na, t) = + let inj = EConstr.Unsafe.to_constr in + LocalAssum (na, inj t) + +let local_def (na, b, t) = + let inj = EConstr.Unsafe.to_constr in + LocalDef (na, inj b, inj t) + let push_rec_types pfix env = let (i, c, t) = pfix in let inj c = EConstr.Unsafe.to_constr c in @@ -31,57 +42,57 @@ let meta_type evd mv = with Not_found -> anomaly (str "unknown meta ?" ++ str (Nameops.string_of_meta mv)) in meta_instance evd ty -let constant_type_knowing_parameters env cst jl = - let paramstyp = Array.map (fun j -> lazy j.uj_type) jl in +let constant_type_knowing_parameters env sigma cst jl = + let paramstyp = Array.map (fun j -> lazy (EConstr.to_constr sigma j.uj_type)) jl in EConstr.of_constr (type_of_constant_knowing_parameters_in env cst paramstyp) -let inductive_type_knowing_parameters env (ind,u) jl = +let inductive_type_knowing_parameters env sigma (ind,u) jl = let mspec = lookup_mind_specif env ind in - let paramstyp = Array.map (fun j -> lazy j.uj_type) jl in + let paramstyp = Array.map (fun j -> lazy (EConstr.to_constr sigma j.uj_type)) jl in EConstr.of_constr (Inductive.type_of_inductive_knowing_parameters env (mspec,u) paramstyp) let e_type_judgment env evdref j = - match EConstr.kind !evdref (EConstr.of_constr (whd_all env !evdref (EConstr.of_constr j.uj_type))) with + match EConstr.kind !evdref (EConstr.of_constr (whd_all env !evdref j.uj_type)) with | Sort s -> {utj_val = j.uj_val; utj_type = s } | Evar ev -> let (evd,s) = Evardefine.define_evar_as_sort env !evdref ev in evdref := evd; { utj_val = j.uj_val; utj_type = s } - | _ -> error_not_type env j + | _ -> error_not_a_type env !evdref j let e_assumption_of_judgment env evdref j = - try EConstr.of_constr (e_type_judgment env evdref j).utj_val - with TypeError _ -> - error_assumption env j + try (e_type_judgment env evdref j).utj_val + with Type_errors.TypeError _ | PretypeError _ -> + error_assumption env !evdref j let e_judge_of_apply env evdref funj argjv = let open EConstr in let rec apply_rec n typ = function | [] -> - { uj_val = Constr.mkApp (j_val funj, Array.map j_val argjv); - uj_type = EConstr.Unsafe.to_constr typ } + { uj_val = mkApp (j_val funj, Array.map j_val argjv); + uj_type = typ } | hj::restjl -> match EConstr.kind !evdref (EConstr.of_constr (whd_all env !evdref typ)) with | Prod (_,c1,c2) -> - if Evarconv.e_cumul env evdref (EConstr.of_constr hj.uj_type) c1 then - apply_rec (n+1) (Vars.subst1 (EConstr.of_constr hj.uj_val) c2) restjl + if Evarconv.e_cumul env evdref hj.uj_type c1 then + apply_rec (n+1) (subst1 hj.uj_val c2) restjl else - error_cant_apply_bad_type env (n, EConstr.Unsafe.to_constr c1, hj.uj_type) funj argjv + error_cant_apply_bad_type env !evdref (n, c1, hj.uj_type) funj argjv | Evar ev -> let (evd',t) = Evardefine.define_evar_as_product !evdref ev in evdref := evd'; let (_,_,c2) = destProd evd' t in - apply_rec (n+1) (Vars.subst1 (EConstr.of_constr hj.uj_val) c2) restjl + apply_rec (n+1) (subst1 hj.uj_val c2) restjl | _ -> - error_cant_apply_not_functional env funj argjv + error_cant_apply_not_functional env !evdref funj argjv in - apply_rec 1 (EConstr.of_constr funj.uj_type) (Array.to_list argjv) + apply_rec 1 funj.uj_type (Array.to_list argjv) let e_check_branch_types env evdref (ind,u) cj (lfj,explft) = if not (Int.equal (Array.length lfj) (Array.length explft)) then - error_number_branches env cj (Array.length explft); + error_number_branches env !evdref cj (Array.length explft); for i = 0 to Array.length explft - 1 do - if not (Evarconv.e_cumul env evdref (EConstr.of_constr lfj.(i).uj_type) (EConstr.of_constr explft.(i))) then - error_ill_formed_branch env cj.uj_val ((ind,i+1),u) lfj.(i).uj_type explft.(i) + if not (Evarconv.e_cumul env evdref lfj.(i).uj_type explft.(i)) then + error_ill_formed_branch env !evdref cj.uj_val ((ind,i+1),u) lfj.(i).uj_type explft.(i) done let max_sort l = @@ -92,13 +103,13 @@ let e_is_correct_arity env evdref c pj ind specif params = let open EConstr in let arsign = make_arity_signature env true (make_ind_family (ind,params)) in let allowed_sorts = elim_sorts specif in - let error () = error_elim_arity env ind allowed_sorts c pj None in + let error () = Pretype_errors.error_elim_arity env !evdref ind allowed_sorts c pj None in let rec srec env pt ar = let pt' = EConstr.of_constr (whd_all env !evdref pt) in match EConstr.kind !evdref pt', ar with | Prod (na1,a1,t), (LocalAssum (_,a1'))::ar' -> if not (Evarconv.e_cumul env evdref a1 (EConstr.of_constr a1')) then error (); - srec (push_rel (LocalAssum (na1,EConstr.Unsafe.to_constr a1)) env) t ar' + srec (push_rel (local_assum (na1,a1)) env) t ar' | Sort s, [] -> if not (Sorts.List.mem (Sorts.family s) allowed_sorts) then error () @@ -106,27 +117,43 @@ let e_is_correct_arity env evdref c pj ind specif params = let evd, s = Evd.fresh_sort_in_family env !evdref (max_sort allowed_sorts) in evdref := Evd.define ev (Constr.mkSort s) evd | _, (LocalDef _ as d)::ar' -> - srec (push_rel d env) (Vars.lift 1 pt') ar' + srec (push_rel d env) (lift 1 pt') ar' | _ -> error () in - srec env (EConstr.of_constr pj.uj_type) (List.rev arsign) + srec env pj.uj_type (List.rev arsign) + +let lambda_applist_assum sigma n c l = + let open EConstr in + let rec app n subst t l = + if Int.equal n 0 then + if l == [] then substl subst t + else anomaly (Pp.str "Not enough arguments") + else match EConstr.kind sigma t, l with + | Lambda(_,_,c), arg::l -> app (n-1) (arg::subst) c l + | LetIn(_,b,_,c), _ -> app (n-1) (substl subst b::subst) c l + | _ -> anomaly (Pp.str "Not enough lambda/let's") in + app n [] c l let e_type_case_branches env evdref (ind,largs) pj c = let specif = lookup_mind_specif env (fst ind) in let nparams = inductive_params specif in let (params,realargs) = List.chop nparams largs in let p = pj.uj_val in + let realargs = List.map EConstr.of_constr realargs in let () = e_is_correct_arity env evdref c pj ind specif params in - let lc = build_branches_type ind specif params p in + let lc = build_branches_type ind specif params (EConstr.to_constr !evdref p) in + let lc = Array.map EConstr.of_constr lc in let n = (snd specif).Declarations.mind_nrealdecls in - let ty = whd_betaiota !evdref (EConstr.of_constr (lambda_applist_assum (n+1) p (realargs@[c]))) in + let ty = whd_betaiota !evdref (lambda_applist_assum !evdref (n+1) p (realargs@[c])) in + let ty = EConstr.of_constr ty in (lc, ty) let e_judge_of_case env evdref ci pj cj lfj = + let open EConstr in let indspec = - try find_mrectype env !evdref (EConstr.of_constr cj.uj_type) - with Not_found -> error_case_not_inductive env cj in + try find_mrectype env !evdref cj.uj_type + with Not_found -> error_case_not_inductive env !evdref cj in let _ = check_case_info env (fst indspec) ci in let (bty,rslty) = e_type_case_branches env evdref indspec pj cj.uj_val in e_check_branch_types env evdref (fst indspec) cj (lfj,bty); @@ -138,27 +165,28 @@ let check_type_fixpoint loc env evdref lna lar vdefj = let lt = Array.length vdefj in if Int.equal (Array.length lar) lt then for i = 0 to lt-1 do - if not (Evarconv.e_cumul env evdref (EConstr.of_constr (vdefj.(i)).uj_type) - (Vars.lift lt lar.(i))) then - Pretype_errors.error_ill_typed_rec_body ~loc env !evdref - i lna vdefj (Array.map EConstr.Unsafe.to_constr lar) + if not (Evarconv.e_cumul env evdref (vdefj.(i)).uj_type + (lift lt lar.(i))) then + error_ill_typed_rec_body ~loc env !evdref + i lna vdefj lar done (* FIXME: might depend on the level of actual parameters!*) let check_allowed_sort env sigma ind c p = let pj = Retyping.get_judgment_of env sigma p in - let ksort = family_of_sort (sort_of_arity env sigma (EConstr.of_constr pj.uj_type)) in + let ksort = family_of_sort (sort_of_arity env sigma pj.uj_type) in let specif = Global.lookup_inductive (fst ind) in let sorts = elim_sorts specif in if not (List.exists ((==) ksort) sorts) then let s = inductive_sort_family (snd specif) in - error_elim_arity env ind sorts (EConstr.Unsafe.to_constr c) pj - (Some(ksort,s,error_elim_explain ksort s)) + error_elim_arity env sigma ind sorts c pj + (Some(ksort,s,Type_errors.error_elim_explain ksort s)) let e_judge_of_cast env evdref cj k tj = + let open EConstr in let expected_type = tj.utj_val in - if not (Evarconv.e_cumul env evdref (EConstr.of_constr cj.uj_type) (EConstr.of_constr expected_type)) then - error_actual_type env cj expected_type; + if not (Evarconv.e_cumul env evdref cj.uj_type expected_type) then + error_actual_type_core env !evdref cj expected_type; { uj_val = mkCast (cj.uj_val, k, expected_type); uj_type = expected_type } @@ -178,11 +206,56 @@ let check_cofix env sigma pcofix = let (idx, (ids, cs, ts)) = pcofix in check_cofix env (idx, (ids, Array.map inj cs, Array.map inj ts)) -let make_judge c ty = - make_judge (EConstr.Unsafe.to_constr c) (EConstr.Unsafe.to_constr ty) - (* The typing machine with universes and existential variables. *) +let judge_of_prop = + { uj_val = EConstr.mkProp; + uj_type = EConstr.mkSort type1_sort } + +let judge_of_set = + { uj_val = EConstr.mkSet; + uj_type = EConstr.mkSort type1_sort } + +let judge_of_prop_contents = function + | Null -> judge_of_prop + | Pos -> judge_of_set + +let judge_of_type u = + let uu = Univ.Universe.super u in + { uj_val = EConstr.mkType u; + uj_type = EConstr.mkType uu } + +let judge_of_relative env v = + Termops.on_judgment EConstr.of_constr (judge_of_relative env v) + +let judge_of_variable env id = + Termops.on_judgment EConstr.of_constr (judge_of_variable env id) + +let judge_of_projection env sigma p cj = + let pb = lookup_projection p env in + let (ind,u), args = + try find_mrectype env sigma cj.uj_type + with Not_found -> error_case_not_inductive env sigma cj + in + let args = List.map EConstr.of_constr args in + let ty = EConstr.of_constr (CVars.subst_instance_constr u pb.Declarations.proj_type) in + let ty = substl (cj.uj_val :: List.rev args) ty in + {uj_val = EConstr.mkProj (p,cj.uj_val); + uj_type = ty} + +let judge_of_abstraction env name var j = + { uj_val = mkLambda (name, var.utj_val, j.uj_val); + uj_type = mkProd (name, var.utj_val, j.uj_type) } + +let judge_of_product env name t1 t2 = + let s = sort_of_product env t1.utj_type t2.utj_type in + { uj_val = mkProd (name, t1.utj_val, t2.utj_val); + uj_type = mkSort s } + +let judge_of_letin env name defj typj j = + { uj_val = mkLetIn (name, defj.uj_val, typj.utj_val, j.uj_val) ; + uj_type = subst1 defj.uj_val j.uj_type } + (* cstr must be in n.f. w.r.t. evars and execute returns a judgement where both the term and type are in n.f. *) let rec execute env evdref cstr = @@ -190,13 +263,13 @@ let rec execute env evdref cstr = let cstr = EConstr.of_constr (whd_evar !evdref (EConstr.Unsafe.to_constr cstr)) in match EConstr.kind !evdref cstr with | Meta n -> - { uj_val = EConstr.Unsafe.to_constr cstr; uj_type = meta_type !evdref n } + { uj_val = cstr; uj_type = EConstr.of_constr (meta_type !evdref n) } | Evar ev -> let ty = EConstr.existential_type !evdref ev in let jty = execute env evdref ty in let jty = e_assumption_of_judgment env evdref jty in - { uj_val = EConstr.Unsafe.to_constr cstr; uj_type = EConstr.Unsafe.to_constr jty } + { uj_val = cstr; uj_type = jty } | Rel n -> judge_of_relative env n @@ -239,7 +312,7 @@ let rec execute env evdref cstr = | Proj (p, c) -> let cj = execute env evdref c in - judge_of_projection env p (Evarutil.j_nf_evar !evdref cj) + judge_of_projection env !evdref p cj | App (f,args) -> let jl = execute_array env evdref args in @@ -248,13 +321,11 @@ let rec execute env evdref cstr = | Ind ind when Environ.template_polymorphic_pind ind env -> (* Sort-polymorphism of inductive types *) make_judge f - (inductive_type_knowing_parameters env ind - (Evarutil.jv_nf_evar !evdref jl)) + (inductive_type_knowing_parameters env !evdref ind jl) | Const cst when Environ.template_polymorphic_pconstant cst env -> (* Sort-polymorphism of inductive types *) make_judge f - (constant_type_knowing_parameters env cst - (Evarutil.jv_nf_evar !evdref jl)) + (constant_type_knowing_parameters env !evdref cst jl) | _ -> execute env evdref f in @@ -263,14 +334,14 @@ let rec execute env evdref cstr = | Lambda (name,c1,c2) -> let j = execute env evdref c1 in let var = e_type_judgment env evdref j in - let env1 = push_rel (LocalAssum (name, var.utj_val)) env in + let env1 = push_rel (local_assum (name, var.utj_val)) env in let j' = execute env1 evdref c2 in judge_of_abstraction env1 name var j' | Prod (name,c1,c2) -> let j = execute env evdref c1 in let varj = e_type_judgment env evdref j in - let env1 = push_rel (LocalAssum (name, varj.utj_val)) env in + let env1 = push_rel (local_assum (name, varj.utj_val)) env in let j' = execute env1 evdref c2 in let varj' = e_type_judgment env1 evdref j' in judge_of_product env name varj varj' @@ -280,7 +351,7 @@ let rec execute env evdref cstr = let j2 = execute env evdref c2 in let j2 = e_type_judgment env evdref j2 in let _ = e_judge_of_cast env evdref j1 DEFAULTcast j2 in - let env1 = push_rel (LocalDef (name, j1.uj_val, j2.utj_val)) env in + let env1 = push_rel (local_def (name, j1.uj_val, j2.utj_val)) env in let j3 = execute env1 evdref c3 in judge_of_letin env name j1 j2 j3 @@ -295,7 +366,7 @@ and execute_recdef env evdref (names,lar,vdef) = let lara = Array.map (e_assumption_of_judgment env evdref) larj in let env1 = push_rec_types (names,lara,vdef) env in let vdefj = execute_array env1 evdref vdef in - let vdefv = Array.map (j_val %> EConstr.of_constr) vdefj in + let vdefv = Array.map j_val vdefj in let _ = check_type_fixpoint Loc.ghost env1 evdref names lara vdefj in (names,lara,vdefv) @@ -304,8 +375,8 @@ and execute_array env evdref = Array.map (execute env evdref) let e_check env evdref c t = let env = enrich_env env evdref in let j = execute env evdref c in - if not (Evarconv.e_cumul env evdref (EConstr.of_constr j.uj_type) t) then - error_actual_type env j (EConstr.to_constr !evdref t) + if not (Evarconv.e_cumul env evdref j.uj_type t) then + error_actual_type_core env !evdref j t (* Type of a constr *) @@ -313,7 +384,7 @@ let unsafe_type_of env evd c = let evdref = ref evd in let env = enrich_env env evdref in let j = execute env evdref c in - j.uj_type + EConstr.Unsafe.to_constr j.uj_type (* Sort of a type *) @@ -331,23 +402,23 @@ let type_of ?(refresh=false) env evd c = let j = execute env evdref c in (* side-effect on evdref *) if refresh then - Evarsolve.refresh_universes ~onlyalg:true (Some false) env !evdref (EConstr.of_constr j.uj_type) - else !evdref, j.uj_type + Evarsolve.refresh_universes ~onlyalg:true (Some false) env !evdref j.uj_type + else !evdref, EConstr.Unsafe.to_constr j.uj_type let e_type_of ?(refresh=false) env evdref c = let env = enrich_env env evdref in let j = execute env evdref c in (* side-effect on evdref *) if refresh then - let evd, c = Evarsolve.refresh_universes ~onlyalg:true (Some false) env !evdref (EConstr.of_constr j.uj_type) in + let evd, c = Evarsolve.refresh_universes ~onlyalg:true (Some false) env !evdref j.uj_type in let () = evdref := evd in c - else j.uj_type + else EConstr.Unsafe.to_constr j.uj_type let e_solve_evars env evdref c = let env = enrich_env env evdref in let c = (execute env evdref c).uj_val in (* side-effect on evdref *) - nf_evar !evdref c + nf_evar !evdref (EConstr.Unsafe.to_constr c) let _ = Evarconv.set_solve_evars (fun env evdref c -> EConstr.of_constr (e_solve_evars env evdref c)) diff --git a/pretyping/typing.mli b/pretyping/typing.mli index 3c1c4324d..1fb414906 100644 --- a/pretyping/typing.mli +++ b/pretyping/typing.mli @@ -6,6 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open Names open Term open Environ open Evd @@ -43,4 +44,11 @@ val check_allowed_sort : env -> evar_map -> pinductive -> EConstr.constr -> ECon (** Raise an error message if bodies have types not unifiable with the expected ones *) val check_type_fixpoint : Loc.t -> env -> evar_map ref -> - Names.Name.t array -> EConstr.types array -> unsafe_judgment array -> unit + Names.Name.t array -> EConstr.types array -> EConstr.unsafe_judgment array -> unit + +val judge_of_prop : EConstr.unsafe_judgment +val judge_of_set : EConstr.unsafe_judgment +val judge_of_abstraction : Environ.env -> Name.t -> + EConstr.unsafe_type_judgment -> EConstr.unsafe_judgment -> EConstr.unsafe_judgment +val judge_of_product : Environ.env -> Name.t -> + EConstr.unsafe_type_judgment -> EConstr.unsafe_type_judgment -> EConstr.unsafe_judgment diff --git a/pretyping/unification.ml b/pretyping/unification.ml index c5c19b49b..1b209fa77 100644 --- a/pretyping/unification.ml +++ b/pretyping/unification.ml @@ -120,6 +120,9 @@ let abstract_list_all env evd typ c l = | UserError _ -> error_cannot_find_well_typed_abstraction env evd p l None | Type_errors.TypeError (env',x) -> + (** FIXME: plug back the typing information *) + error_cannot_find_well_typed_abstraction env evd p l None + | Pretype_errors.PretypeError (env',evd,TypingError x) -> error_cannot_find_well_typed_abstraction env evd p l (Some (env',x)) in let typp = EConstr.of_constr typp in evd,(p,typp) @@ -1255,15 +1258,12 @@ let is_mimick_head sigma ts f = | (Rel _|Construct _|Ind _) -> true | _ -> false -let make_judge c t = - make_judge (EConstr.Unsafe.to_constr c) (EConstr.Unsafe.to_constr t) - let try_to_coerce env evd c cty tycon = let j = make_judge c cty in let (evd',j') = inh_conv_coerce_rigid_to true Loc.ghost env evd j tycon in let evd' = Evarconv.consider_remaining_unif_problems env evd' in let evd' = Evd.map_metas_fvalue (nf_evar evd') evd' in - (evd',EConstr.of_constr j'.uj_val) + (evd',j'.uj_val) let w_coerce_to_type env evd c cty mvty = let evd,tycon = pose_all_metas_as_evars env evd mvty in diff --git a/proofs/clenv.ml b/proofs/clenv.ml index 2d621e97c..956be88c8 100644 --- a/proofs/clenv.ml +++ b/proofs/clenv.ml @@ -669,12 +669,14 @@ let evar_of_binder holes = function user_err (str "No such binder.") let define_with_type sigma env ev c = + let c = EConstr.of_constr c in let t = Retyping.get_type_of env sigma (EConstr.of_constr ev) in - let ty = Retyping.get_type_of env sigma (EConstr.of_constr c) in + let ty = Retyping.get_type_of env sigma c in + let ty = EConstr.of_constr ty in let j = Environ.make_judge c ty in let (sigma, j) = Coercion.inh_conv_coerce_to true (Loc.ghost) env sigma j (EConstr.of_constr t) in let (ev, _) = destEvar ev in - let sigma = Evd.define ev j.Environ.uj_val sigma in + let sigma = Evd.define ev (EConstr.Unsafe.to_constr j.Environ.uj_val) sigma in sigma let solve_evar_clause env sigma hyp_only clause = function diff --git a/proofs/refine.ml b/proofs/refine.ml index 0ed74c9b3..067764c00 100644 --- a/proofs/refine.ml +++ b/proofs/refine.ml @@ -133,7 +133,9 @@ let refine ?(unsafe = true) f = (** Useful definitions *) let with_type env evd c t = - let my_type = Retyping.get_type_of env evd (EConstr.of_constr c) in + let c = EConstr.of_constr c in + let my_type = Retyping.get_type_of env evd c in + let my_type = EConstr.of_constr my_type in let j = Environ.make_judge c my_type in let (evd,j') = Coercion.inh_conv_coerce_to true (Loc.ghost) env evd j (EConstr.of_constr t) @@ -147,6 +149,7 @@ let refine_casted ?unsafe f = Proofview.Goal.enter { enter = begin fun gl -> let f = { run = fun h -> let Sigma (c, h, p) = f.run h in let sigma, c = with_type env (Sigma.to_evar_map h) c concl in + let c = EConstr.Unsafe.to_constr c in Sigma (c, Sigma.Unsafe.of_evar_map sigma, p) } in refine ?unsafe f diff --git a/proofs/refine.mli b/proofs/refine.mli index 3d140f036..4158e446c 100644 --- a/proofs/refine.mli +++ b/proofs/refine.mli @@ -36,7 +36,7 @@ val refine_one : ?unsafe:bool -> ('a * Constr.t) Sigma.run -> 'a tactic (** {7 Helper functions} *) val with_type : Environ.env -> Evd.evar_map -> - Term.constr -> Term.types -> Evd.evar_map * Term.constr + Term.constr -> Term.types -> Evd.evar_map * EConstr.constr (** [with_type env sigma c t] ensures that [c] is of type [t] inserting a coercion if needed. *) diff --git a/tactics/contradiction.ml b/tactics/contradiction.ml index 789028ac1..9580fdbfc 100644 --- a/tactics/contradiction.ml +++ b/tactics/contradiction.ml @@ -30,7 +30,7 @@ let absurd c = let sigma = Sigma.to_evar_map sigma in let j = Retyping.get_judgment_of env sigma (EConstr.of_constr c) in let sigma, j = Coercion.inh_coerce_to_sort Loc.ghost env sigma j in - let t = j.Environ.utj_val in + let t = EConstr.Unsafe.to_constr j.Environ.utj_val in let tac = Tacticals.New.tclTHENLIST [ elim_type (build_coq_False ()); diff --git a/toplevel/explainErr.ml b/toplevel/explainErr.ml index 17897460c..414087718 100644 --- a/toplevel/explainErr.ml +++ b/toplevel/explainErr.ml @@ -63,6 +63,7 @@ let process_vernac_interp_error with_header exn = match fst exn with mt() in wrap_vernac_error with_header exn (str "Universe inconsistency" ++ msg ++ str ".") | TypeError(ctx,te) -> + let te = Himsg.map_ptype_error EConstr.of_constr te in wrap_vernac_error with_header exn (Himsg.explain_type_error ctx Evd.empty te) | PretypeError(ctx,sigma,te) -> wrap_vernac_error with_header exn (Himsg.explain_pretype_error ctx sigma te) diff --git a/toplevel/himsg.ml b/toplevel/himsg.ml index 20cda947a..6e63a71fd 100644 --- a/toplevel/himsg.ml +++ b/toplevel/himsg.ml @@ -83,6 +83,8 @@ let rec contract3' env a b c = function (** Ad-hoc reductions *) +let to_unsafe_judgment j = on_judgment EConstr.Unsafe.to_constr j + let j_nf_betaiotaevar sigma j = { uj_val = Evarutil.nf_evar sigma j.uj_val; uj_type = Reductionops.nf_betaiota sigma (EConstr.of_constr j.uj_type) } @@ -170,6 +172,7 @@ let explain_unbound_var env v = let explain_not_type env sigma j = let j = Evarutil.j_nf_evar sigma j in + let j = to_unsafe_judgment j in let pe = pr_ne_context_of (str "In environment") env sigma in let pc,pt = pr_ljudge_env env sigma j in pe ++ str "The term" ++ brk(1,1) ++ pc ++ spc () ++ @@ -177,6 +180,7 @@ let explain_not_type env sigma j = str "which should be Set, Prop or Type." let explain_bad_assumption env sigma j = + let j = to_unsafe_judgment j in let pe = pr_ne_context_of (str "In environment") env sigma in let pc,pt = pr_ljudge_env env sigma j in pe ++ str "Cannot declare a variable or hypothesis over the term" ++ @@ -184,6 +188,7 @@ let explain_bad_assumption env sigma j = str "because this term is not a type." let explain_reference_variables id c = + let c = EConstr.Unsafe.to_constr c in (* c is intended to be a global reference *) let pc = pr_global (Globnames.global_of_constr c) in pc ++ strbrk " depends on the variable " ++ pr_id id ++ @@ -202,6 +207,8 @@ let pr_puniverses f env (c,u) = else mt()) let explain_elim_arity env sigma ind sorts c pj okinds = + let c = EConstr.Unsafe.to_constr c in + let pj = to_unsafe_judgment pj in let env = make_all_name_different env in let pi = pr_inductive env (fst ind) in let pc = pr_lconstr_env env sigma c in @@ -237,6 +244,7 @@ let explain_elim_arity env sigma ind sorts c pj okinds = let explain_case_not_inductive env sigma cj = let cj = Evarutil.j_nf_evar sigma cj in + let cj = to_unsafe_judgment cj in let env = make_all_name_different env in let pc = pr_lconstr_env env sigma cj.uj_val in let pct = pr_lconstr_env env sigma cj.uj_type in @@ -250,6 +258,7 @@ let explain_case_not_inductive env sigma cj = let explain_number_branches env sigma cj expn = let cj = Evarutil.j_nf_evar sigma cj in + let cj = to_unsafe_judgment cj in let env = make_all_name_different env in let pc = pr_lconstr_env env sigma cj.uj_val in let pct = pr_lconstr_env env sigma cj.uj_type in @@ -258,6 +267,9 @@ let explain_number_branches env sigma cj expn = str "expects " ++ int expn ++ str " branches." let explain_ill_formed_branch env sigma c ci actty expty = + let c = EConstr.Unsafe.to_constr c in + let actty = EConstr.Unsafe.to_constr actty in + let expty = EConstr.Unsafe.to_constr expty in let simp t = Reduction.nf_betaiota env (Evarutil.nf_evar sigma t) in let c = Evarutil.nf_evar sigma c in let env = make_all_name_different env in @@ -270,6 +282,8 @@ let explain_ill_formed_branch env sigma c ci actty expty = str "which should be" ++ brk(1,1) ++ pe ++ str "." let explain_generalization env sigma (name,var) j = + let var = EConstr.Unsafe.to_constr var in + let j = to_unsafe_judgment j in let pe = pr_ne_context_of (str "In environment") env sigma in let pv = pr_ltype_env env sigma var in let (pc,pt) = pr_ljudge_env (push_rel_assum (name,EConstr.of_constr var) env) sigma j in @@ -343,6 +357,8 @@ let explain_unification_error env sigma p1 p2 = function prlist_with_sep pr_semicolon (fun x -> x) l ++ str ")" let explain_actual_type env sigma j t reason = + let j = to_unsafe_judgment j in + let t = EConstr.Unsafe.to_constr t in let env = make_all_name_different env in let j = j_nf_betaiotaevar sigma j in let t = Reductionops.nf_betaiota sigma (EConstr.of_constr t) in @@ -359,10 +375,14 @@ let explain_actual_type env sigma j t reason = ppreason ++ str ".") let explain_cant_apply_bad_type env sigma (n,exptyp,actualtyp) rator randl = + let exptyp = EConstr.Unsafe.to_constr exptyp in + let actualtyp = EConstr.Unsafe.to_constr actualtyp in + let randl = Array.map to_unsafe_judgment randl in let randl = jv_nf_betaiotaevar sigma randl in let exptyp = Evarutil.nf_evar sigma exptyp in let actualtyp = Reductionops.nf_betaiota sigma (EConstr.of_constr actualtyp) in let rator = Evarutil.j_nf_evar sigma rator in + let rator = to_unsafe_judgment rator in let env = make_all_name_different env in let actualtyp, exptyp = pr_explicit env sigma actualtyp exptyp in let nargs = Array.length randl in @@ -388,7 +408,9 @@ let explain_cant_apply_bad_type env sigma (n,exptyp,actualtyp) rator randl = let explain_cant_apply_not_functional env sigma rator randl = let randl = Evarutil.jv_nf_evar sigma randl in + let randl = Array.map to_unsafe_judgment randl in let rator = Evarutil.j_nf_evar sigma rator in + let rator = to_unsafe_judgment rator in let env = make_all_name_different env in let nargs = Array.length randl in (* let pe = pr_ne_context_of (str "in environment") env sigma in*) @@ -424,6 +446,7 @@ let explain_not_product env sigma c = (* TODO: use the names *) (* (co)fixpoints *) let explain_ill_formed_rec_body env sigma err names i fixenv vdefj = + let pr_lconstr_env env sigma c = pr_lconstr_env env sigma (EConstr.Unsafe.to_constr c) in let prt_name i = match names.(i) with Name id -> str "Recursive definition of " ++ pr_id id @@ -511,6 +534,8 @@ let explain_ill_formed_rec_body env sigma err names i fixenv vdefj = let explain_ill_typed_rec_body env sigma i names vdefj vargs = let vdefj = Evarutil.jv_nf_evar sigma vdefj in + let vargs = Array.map EConstr.Unsafe.to_constr vargs in + let vdefj = Array.map to_unsafe_judgment vdefj in let vargs = Array.map (Evarutil.nf_evar sigma) vargs in let env = make_all_name_different env in let pvd = pr_lconstr_env env sigma vdefj.(i).uj_val in @@ -522,6 +547,7 @@ let explain_ill_typed_rec_body env sigma i names vdefj vargs = str "while it should be" ++ spc () ++ pv ++ str "." let explain_cant_find_case_type env sigma c = + let c = EConstr.Unsafe.to_constr c in let c = Evarutil.nf_evar sigma c in let env = make_all_name_different env in let pe = pr_lconstr_env env sigma c in @@ -822,14 +848,18 @@ let explain_pretype_error env sigma err = match err with | CantFindCaseType c -> explain_cant_find_case_type env sigma c | ActualTypeNotCoercible (j,t,e) -> + let j = to_unsafe_judgment j in + let t = EConstr.Unsafe.to_constr t in let {uj_val = c; uj_type = actty} = j in let (env, c, actty, expty), e = contract3' env c actty t e in let j = {uj_val = c; uj_type = actty} in - explain_actual_type env sigma j expty (Some e) + explain_actual_type env sigma (on_judgment EConstr.of_constr j) (EConstr.of_constr expty) (Some e) | UnifOccurCheck (ev,rhs) -> explain_occur_check env sigma ev rhs | UnsolvableImplicit (evk,exp) -> explain_unsolvable_implicit env sigma evk exp | VarNotFound id -> explain_var_not_found env id | UnexpectedType (actual,expect) -> + let actual = EConstr.Unsafe.to_constr actual in + let expect = EConstr.Unsafe.to_constr expect in let env, actual, expect = contract2 env actual expect in explain_unexpected_type env sigma actual expect | NotProduct c -> explain_not_product env sigma c @@ -1289,8 +1319,48 @@ let explain_pattern_matching_error env sigma = function | CannotInferPredicate typs -> explain_cannot_infer_predicate env sigma typs +let map_pguard_error f = function +| NotEnoughAbstractionInFixBody -> NotEnoughAbstractionInFixBody +| RecursionNotOnInductiveType c -> RecursionNotOnInductiveType (f c) +| RecursionOnIllegalTerm (n, (env, c), l1, l2) -> RecursionOnIllegalTerm (n, (env, f c), l1, l2) +| NotEnoughArgumentsForFixCall n -> NotEnoughArgumentsForFixCall n +| CodomainNotInductiveType c -> CodomainNotInductiveType (f c) +| NestedRecursiveOccurrences -> NestedRecursiveOccurrences +| UnguardedRecursiveCall c -> UnguardedRecursiveCall (f c) +| RecCallInTypeOfAbstraction c -> RecCallInTypeOfAbstraction (f c) +| RecCallInNonRecArgOfConstructor c -> RecCallInNonRecArgOfConstructor (f c) +| RecCallInTypeOfDef c -> RecCallInTypeOfDef (f c) +| RecCallInCaseFun c -> RecCallInCaseFun (f c) +| RecCallInCaseArg c -> RecCallInCaseArg (f c) +| RecCallInCasePred c -> RecCallInCasePred (f c) +| NotGuardedForm c -> NotGuardedForm (f c) +| ReturnPredicateNotCoInductive c -> ReturnPredicateNotCoInductive (f c) + +let map_ptype_error f = function +| UnboundRel n -> UnboundRel n +| UnboundVar id -> UnboundVar id +| NotAType j -> NotAType (on_judgment f j) +| BadAssumption j -> BadAssumption (on_judgment f j) +| ReferenceVariables (id, c) -> ReferenceVariables (id, f c) +| ElimArity (pi, dl, c, j, ar) -> ElimArity (pi, dl, f c, on_judgment f j, ar) +| CaseNotInductive j -> CaseNotInductive (on_judgment f j) +| WrongCaseInfo (pi, ci) -> WrongCaseInfo (pi, ci) +| NumberBranches (j, n) -> NumberBranches (on_judgment f j, n) +| IllFormedBranch (c, pc, t1, t2) -> IllFormedBranch (f c, pc, f t1, f t2) +| Generalization ((na, t), j) -> Generalization ((na, f t), on_judgment f j) +| ActualType (j, t) -> ActualType (on_judgment f j, f t) +| CantApplyBadType ((n, c1, c2), j, vj) -> + CantApplyBadType ((n, f c1, f c2), on_judgment f j, Array.map (on_judgment f) vj) +| CantApplyNonFunctional (j, jv) -> CantApplyNonFunctional (on_judgment f j, Array.map (on_judgment f) jv) +| IllFormedRecBody (ge, na, n, env, jv) -> + IllFormedRecBody (map_pguard_error f ge, na, n, env, Array.map (on_judgment f) jv) +| IllTypedRecBody (n, na, jv, t) -> + IllTypedRecBody (n, na, Array.map (on_judgment f) jv, Array.map f t) +| UnsatisfiedConstraints g -> UnsatisfiedConstraints g + let explain_reduction_tactic_error = function | Tacred.InvalidAbstraction (env,sigma,c,(env',e)) -> + let e = map_ptype_error EConstr.of_constr e in let c = EConstr.to_constr sigma c in str "The abstracted term" ++ spc () ++ quote (pr_goal_concl_style_env env sigma c) ++ diff --git a/toplevel/himsg.mli b/toplevel/himsg.mli index ced54fd27..6915ea921 100644 --- a/toplevel/himsg.mli +++ b/toplevel/himsg.mli @@ -40,3 +40,6 @@ val explain_module_error : Modops.module_typing_error -> std_ppcmds val explain_module_internalization_error : Modintern.module_internalization_error -> std_ppcmds + +val map_pguard_error : ('c -> 'd) -> 'c pguard_error -> 'd pguard_error +val map_ptype_error : ('c -> 'd) -> ('c, 'c) ptype_error -> ('d, 'd) ptype_error diff --git a/toplevel/vernacentries.ml b/toplevel/vernacentries.ml index 01b15407b..3b8c62738 100644 --- a/toplevel/vernacentries.ml +++ b/toplevel/vernacentries.ml @@ -1579,7 +1579,8 @@ let vernac_check_may_eval redexp glopt rc = let c = nf c in let j = if Evarutil.has_undefined_evars sigma' (EConstr.of_constr c) then - Evarutil.j_nf_evar sigma' (Retyping.get_judgment_of env sigma' (EConstr.of_constr c)) + let j = Evarutil.j_nf_evar sigma' (Retyping.get_judgment_of env sigma' (EConstr.of_constr c)) in + Termops.on_judgment EConstr.Unsafe.to_constr j else (* OK to call kernel which does not support evars *) Arguments_renaming.rename_typing env c in -- cgit v1.2.3 From 536026f3e20f761e8ef366ed732da7d3b626ac5e Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 11 Nov 2016 15:39:01 +0100 Subject: Cleaning up opening of the EConstr module in pretyping folder. --- plugins/extraction/extraction.ml | 1 + pretyping/cases.ml | 1 - pretyping/coercion.ml | 11 ++--- pretyping/constr_matching.ml | 8 ++-- pretyping/detyping.ml | 1 + pretyping/evarconv.ml | 49 ++++++++----------- pretyping/evardefine.ml | 27 +++++------ pretyping/evarsolve.ml | 79 +++++++++++-------------------- pretyping/patternops.ml | 2 +- pretyping/reductionops.ml | 100 ++++++++++++++++++--------------------- pretyping/retyping.ml | 19 ++++---- pretyping/retyping.mli | 2 +- pretyping/tacred.ml | 47 ++++-------------- pretyping/typing.ml | 7 --- pretyping/unification.ml | 28 +---------- tactics/tactics.ml | 2 +- 16 files changed, 135 insertions(+), 249 deletions(-) diff --git a/plugins/extraction/extraction.ml b/plugins/extraction/extraction.ml index 0c4fa7055..6559aeb08 100644 --- a/plugins/extraction/extraction.ml +++ b/plugins/extraction/extraction.ml @@ -599,6 +599,7 @@ let rec extract_term env mle mlt c args = extract_cons_app env mle mlt cp args | Proj (p, c) -> let term = Retyping.expand_projection env (Evd.from_env env) p (EConstr.of_constr c) [] in + let term = EConstr.Unsafe.to_constr term in extract_term env mle mlt term args | Rel n -> (* As soon as the expected [mlt] for the head is known, *) diff --git a/pretyping/cases.ml b/pretyping/cases.ml index b43e2193a..57d12a19f 100644 --- a/pretyping/cases.ml +++ b/pretyping/cases.ml @@ -111,7 +111,6 @@ let make_anonymous_patvars n = let relocate_rel n1 n2 k j = if Int.equal j (n1 + k) then n2+k else j let rec relocate_index sigma n1 n2 k t = - let open EConstr in match EConstr.kind sigma t with | Rel j when Int.equal j (n1 + k) -> mkRel (n2+k) | Rel j when j < n1+k -> t diff --git a/pretyping/coercion.ml b/pretyping/coercion.ml index 2d4296fe4..e7279df7a 100644 --- a/pretyping/coercion.ml +++ b/pretyping/coercion.ml @@ -18,6 +18,7 @@ open CErrors open Util open Names open Term +open EConstr open Vars open Reductionops open Environ @@ -48,7 +49,6 @@ exception NoCoercionNoUnifier of evar_map * unification_error (* Here, funj is a coercion therefore already typed in global context *) let apply_coercion_args env evd check isproj argl funj = - let open EConstr in let evdref = ref evd in let rec apply_rec acc typ = function | [] -> @@ -68,7 +68,7 @@ let apply_coercion_args env evd check isproj argl funj = | Prod (_,c1,c2) -> if check && not (e_cumul env evdref (EConstr.of_constr (Retyping.get_type_of env !evdref h)) c1) then raise NoCoercion; - apply_rec (h::acc) (Vars.subst1 h c2) restl + apply_rec (h::acc) (subst1 h c2) restl | _ -> anomaly (Pp.str "apply_coercion_args") in let res = apply_rec [] funj.uj_type argl in @@ -121,9 +121,8 @@ let hnf env evd c = whd_all env evd c let hnf_nodelta env evd c = whd_betaiota evd c let lift_args n sign = - let open EConstr in let rec liftrec k = function - | t::sign -> Vars.liftn n k t :: (liftrec (k-1) sign) + | t::sign -> liftn n k t :: (liftrec (k-1) sign) | [] -> [] in liftrec (List.length sign) sign @@ -150,8 +149,6 @@ let mu env evdref t = and coerce loc env evdref (x : EConstr.constr) (y : EConstr.constr) : (EConstr.constr -> EConstr.constr) option = - let open EConstr in - let open Vars in let open Context.Rel.Declaration in let rec coerce_unify env x y = let x = hnf env !evdref x and y = hnf env !evdref y in @@ -478,8 +475,6 @@ let inh_coerce_to_fail env evd rigidonly v t c1 = with UnableToUnify _ -> raise NoCoercion let rec inh_conv_coerce_to_fail loc env evd rigidonly v t c1 = - let open EConstr in - let open Vars in try (the_conv_x_leq env t c1 evd, v) with UnableToUnify (best_failed_evd,e) -> try inh_coerce_to_fail env evd rigidonly v t c1 diff --git a/pretyping/constr_matching.ml b/pretyping/constr_matching.ml index ecf6b1121..4d2500ccd 100644 --- a/pretyping/constr_matching.ml +++ b/pretyping/constr_matching.ml @@ -238,7 +238,7 @@ let matches_core env sigma convert allow_partial_app allow_bound_rels match EConstr.kind sigma c2 with | Proj (pr, c) when not (Projection.unfolded pr) -> (try let term = Retyping.expand_projection env sigma pr c (Array.to_list args2) in - sorec ctx env subst p (EConstr.of_constr term) + sorec ctx env subst p term with Retyping.RetypeError _ -> raise PatternMatchingFailure) | _ -> raise PatternMatchingFailure) @@ -254,7 +254,7 @@ let matches_core env sigma convert allow_partial_app allow_bound_rels else raise PatternMatchingFailure | _, Proj (pr,c) when not (Projection.unfolded pr) -> (try let term = Retyping.expand_projection env sigma pr c (Array.to_list arg2) in - sorec ctx env subst p (EConstr.of_constr term) + sorec ctx env subst p term with Retyping.RetypeError _ -> raise PatternMatchingFailure) | _, _ -> try Array.fold_left2 (sorec ctx env) (sorec ctx env subst c1 c2) arg1 arg2 @@ -266,7 +266,7 @@ let matches_core env sigma convert allow_partial_app allow_bound_rels | PApp (c, args), Proj (pr, c2) -> (try let term = Retyping.expand_projection env sigma pr c2 [] in - sorec ctx env subst p (EConstr.of_constr term) + sorec ctx env subst p term with Retyping.RetypeError _ -> raise PatternMatchingFailure) | PProj (p1,c1), Proj (p2,c2) when Projection.equal p1 p2 -> @@ -460,7 +460,7 @@ let sub_match ?(partial_app=false) ?(closed=true) env sigma pat c = if partial_app then try let term = Retyping.expand_projection env sigma p c' [] in - aux env (EConstr.of_constr term) mk_ctx next + aux env term mk_ctx next with Retyping.RetypeError _ -> next () else try_aux [env, c'] next_mk_ctx next diff --git a/pretyping/detyping.ml b/pretyping/detyping.ml index e5e778f23..4756ec30e 100644 --- a/pretyping/detyping.ml +++ b/pretyping/detyping.ml @@ -513,6 +513,7 @@ let rec detype flags avoid env sigma t = if print_primproj_params () then try let c = Retyping.expand_projection (snd env) sigma p (EConstr.of_constr c) [] in + let c = EConstr.Unsafe.to_constr c in detype flags avoid env sigma c with Retyping.RetypeError _ -> noparams () else noparams () diff --git a/pretyping/evarconv.ml b/pretyping/evarconv.ml index 639d6260e..77e91095f 100644 --- a/pretyping/evarconv.ml +++ b/pretyping/evarconv.ml @@ -6,15 +6,18 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +module CVars = Vars + open CErrors open Util open Names open Term +open Termops +open EConstr open Vars open CClosure open Reduction open Reductionops -open Termops open Environ open Recordops open Evarutil @@ -43,14 +46,12 @@ let _ = Goptions.declare_bool_option { } let unfold_projection env evd ts p c = - let open EConstr in let cst = Projection.constant p in if is_transparent_constant ts cst then Some (mkProj (Projection.make cst true, c)) else None let eval_flexible_term ts env evd c = - let open EConstr in match EConstr.kind evd c with | Const (c,u as cu) -> if is_transparent_constant ts c @@ -59,7 +60,7 @@ let eval_flexible_term ts env evd c = | Rel n -> (try match lookup_rel n env with | RelDecl.LocalAssum _ -> None - | RelDecl.LocalDef (_,v,_) -> Some (Vars.lift n (EConstr.of_constr v)) + | RelDecl.LocalDef (_,v,_) -> Some (lift n (EConstr.of_constr v)) with Not_found -> None) | Var id -> (try @@ -67,7 +68,7 @@ let eval_flexible_term ts env evd c = Option.map EConstr.of_constr (env |> lookup_named id |> NamedDecl.get_value) else None with Not_found -> None) - | LetIn (_,b,_,c) -> Some (Vars.subst1 b c) + | LetIn (_,b,_,c) -> Some (subst1 b c) | Lambda _ -> Some c | Proj (p, c) -> if Projection.unfolded p then assert false @@ -105,7 +106,6 @@ let position_problem l2r = function | CUMUL -> Some l2r let occur_rigidly (evk,_ as ev) evd t = - let open EConstr in let rec aux t = match EConstr.kind evd t with | App (f, c) -> if aux f then Array.exists aux c else false @@ -141,14 +141,13 @@ let occur_rigidly (evk,_ as ev) evd t = projection would have been reduced) *) let check_conv_record env sigma (t1,sk1) (t2,sk2) = - let open EConstr in let (proji, u), arg = Termops.global_app_of_constr sigma t1 in let canon_s,sk2_effective = try match EConstr.kind sigma t2 with Prod (_,a,b) -> (* assert (l2=[]); *) let _, a, b = destProd sigma t2 in - if Vars.noccurn sigma 1 b then + if noccurn sigma 1 b then lookup_canonical_conversion (proji, Prod_cs), (Stack.append_app [|a;EConstr.of_constr (pop b)|] Stack.empty) else raise Not_found @@ -185,9 +184,9 @@ let check_conv_record env sigma (t1,sk1) (t2,sk2) = | None -> raise Not_found | Some (l',el,s') -> (l'@Stack.append_app [|el|] Stack.empty,s') in let subst, ctx' = Universes.fresh_universe_context_set_instance ctx in - let c' = EConstr.of_constr (subst_univs_level_constr subst c) in - let t' = subst_univs_level_constr subst t' in - let bs' = List.map (subst_univs_level_constr subst %> EConstr.of_constr) bs in + let c' = EConstr.of_constr (CVars.subst_univs_level_constr subst c) in + let t' = CVars.subst_univs_level_constr subst t' in + let bs' = List.map (CVars.subst_univs_level_constr subst %> EConstr.of_constr) bs in let h, _ = decompose_app_vect sigma (EConstr.of_constr t') in ctx',(EConstr.of_constr h, t2),c',bs',(Stack.append_app_list params Stack.empty,params1), (Stack.append_app_list us Stack.empty,us2),(extra_args1,extra_args2),c1, @@ -379,7 +378,6 @@ let rec evar_conv_x ts env evd pbty term1 term2 = and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty ((term1,sk1 as appr1),csts1) ((term2,sk2 as appr2),csts2) = - let open EConstr in let quick_fail i = (* not costly, loses info *) UnifFailure (i, NotSameHead) in @@ -466,7 +464,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty let termM' = Retyping.expand_projection env evd p c [] in let apprM', cstsM' = whd_betaiota_deltazeta_for_iota_state - (fst ts) env evd cstsM (EConstr.of_constr termM',skM) + (fst ts) env evd cstsM (termM',skM) in let delta' i = switch (evar_eqappr_x ts env i pbty) (apprF,cstsF) (apprM',cstsM') @@ -642,7 +640,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty (* Catch the p.c ~= p c' cases *) | Proj (p,c), Const (p',u) when eq_constant (Projection.constant p) p' -> let res = - try Some (destApp evd (EConstr.of_constr (Retyping.expand_projection env evd p c []))) + try Some (destApp evd (Retyping.expand_projection env evd p c [])) with Retyping.RetypeError _ -> None in (match res with @@ -653,7 +651,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty | Const (p,u), Proj (p',c') when eq_constant p (Projection.constant p') -> let res = - try Some (destApp evd (EConstr.of_constr (Retyping.expand_projection env evd p' c' []))) + try Some (destApp evd (Retyping.expand_projection env evd p' c' [])) with Retyping.RetypeError _ -> None in (match res with @@ -699,7 +697,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty | Lambda _ -> assert (match args with [] -> true | _ -> false); true | LetIn (_,b,_,c) -> is_unnamed (fst (whd_betaiota_deltazeta_for_iota_state - (fst ts) env i Cst_stack.empty (EConstr.Vars.subst1 b c, args))) + (fst ts) env i Cst_stack.empty (subst1 b c, args))) | Fix _ -> true (* Partially applied fix can be the result of a whd call *) | Proj (p, _) -> Projection.unfolded p || Stack.not_purely_applicative args | Case _ | App _| Cast _ -> assert false in @@ -878,7 +876,6 @@ and conv_record trs env evd (ctx,(h,h2),c,bs,(params,params1),(us,us2),(sk1,sk2) had to be initially resolved *) - let open EConstr in let evd = Evd.merge_context_set Evd.univ_flexible evd ctx in if Reductionops.Stack.compare_shape sk1 sk2 then let (evd',ks,_,test) = @@ -886,12 +883,12 @@ and conv_record trs env evd (ctx,(h,h2),c,bs,(params,params1),(us,us2),(sk1,sk2) (fun (i,ks,m,test) b -> if match n with Some n -> Int.equal m n | None -> false then let ty = EConstr.of_constr (Retyping.get_type_of env i t2) in - let test i = evar_conv_x trs env i CUMUL ty (Vars.substl ks b) in + let test i = evar_conv_x trs env i CUMUL ty (substl ks b) in (i,t2::ks, m-1, test) else let dloc = (Loc.ghost,Evar_kinds.InternalHole) in let i = Sigma.Unsafe.of_evar_map i in - let Sigma (ev, i', _) = Evarutil.new_evar env i ~src:dloc (Vars.substl ks b) in + let Sigma (ev, i', _) = Evarutil.new_evar env i ~src:dloc (substl ks b) in let i' = Sigma.to_evar_map i' in (i', EConstr.of_constr ev :: ks, m - 1,test)) (evd,[],List.length bs,fun i -> Success i) bs @@ -900,17 +897,17 @@ and conv_record trs env evd (ctx,(h,h2),c,bs,(params,params1),(us,us2),(sk1,sk2) ise_and evd' [(fun i -> exact_ise_stack2 env i - (fun env' i' cpb x1 x -> evar_conv_x trs env' i' cpb x1 (Vars.substl ks x)) + (fun env' i' cpb x1 x -> evar_conv_x trs env' i' cpb x1 (substl ks x)) params1 params); (fun i -> exact_ise_stack2 env i - (fun env' i' cpb u1 u -> evar_conv_x trs env' i' cpb u1 (Vars.substl ks u)) + (fun env' i' cpb u1 u -> evar_conv_x trs env' i' cpb u1 (substl ks u)) us2 us); (fun i -> evar_conv_x trs env i CONV c1 app); (fun i -> exact_ise_stack2 env i (evar_conv_x trs) sk1 sk2); test; (fun i -> evar_conv_x trs env i CONV h2 - (EConstr.of_constr (fst (decompose_app_vect i (Vars.substl ks h)))))] + (EConstr.of_constr (fst (decompose_app_vect i (substl ks h)))))] else UnifFailure(evd,(*dummy*)NotSameHead) and eta_constructor ts env evd sk1 ((ind, i), u) sk2 term2 = @@ -951,7 +948,6 @@ let set_evar_conv f = Hook.set evar_conv_hook_set f (* We assume here |l1| <= |l2| *) let first_order_unification ts env evd (ev1,l1) (term2,l2) = - let open EConstr in let (deb2,rest2) = Array.chop (Array.length l2-Array.length l1) l2 in ise_and evd (* First compare extra args for better failure message *) @@ -973,7 +969,6 @@ let choose_less_dependent_instance evk evd term args = | (id, _) :: _ -> Some (Evd.define evk (Constr.mkVar id) evd) let apply_on_subterm env evdref f c t = - let open EConstr in let rec applyrec (env,(k,c) as acc) t = (* By using eq_constr, we make an approximation, for instance, we *) (* could also be interested in finding a term u convertible to t *) @@ -987,7 +982,7 @@ let apply_on_subterm env evdref f c t = mkEvar (evk, Array.of_list (List.map2 g ctx (Array.to_list args))) | _ -> map_constr_with_binders_left_to_right !evdref - (fun d (env,(k,c)) -> (push_rel d env, (k+1,Vars.lift 1 c))) + (fun d (env,(k,c)) -> (push_rel d env, (k+1,lift 1 c))) applyrec acc t in applyrec (env,(0,c)) t @@ -996,7 +991,6 @@ let filter_possible_projections evd c ty ctxt args = (* Since args in the types will be replaced by holes, we count the fv of args to have a well-typed filter; don't know how necessary it is however to have a well-typed filter here *) - let open EConstr in let fv1 = free_rels evd (mkApp (c,args)) (* Hack: locally untyped *) in let fv2 = collect_vars evd (mkApp (c,args)) in let len = Array.length args in @@ -1039,7 +1033,6 @@ let set_solve_evars f = solve_evars := f exception TypingFailed of evar_map let second_order_matching ts env_rhs evd (evk,args) argoccs rhs = - let open EConstr in try let evi = Evd.find_undefined evd evk in let env_evar = evar_filtered_env evi in @@ -1137,7 +1130,6 @@ let to_pb (pb, env, t1, t2) = (pb, env, EConstr.Unsafe.to_constr t1, EConstr.Unsafe.to_constr t2) let second_order_matching_with_args ts env evd pbty ev l t = - let open EConstr in (* let evd,ev = evar_absorb_arguments env evd ev l in let argoccs = Array.map_to_list (fun _ -> None) (snd ev) in @@ -1150,7 +1142,6 @@ let second_order_matching_with_args ts env evd pbty ev l t = UnifFailure (evd, CannotSolveConstraint (pb,ProblemBeyondCapabilities)) let apply_conversion_problem_heuristic ts env evd pbty t1 t2 = - let open EConstr in let t1 = apprec_nohdbeta ts env evd (whd_head_evar evd t1) in let t2 = apprec_nohdbeta ts env evd (whd_head_evar evd t2) in let (term1,l1 as appr1) = try destApp evd t1 with DestKO -> (t1, [||]) in diff --git a/pretyping/evardefine.ml b/pretyping/evardefine.ml index ff40a6938..fa3b9ca0b 100644 --- a/pretyping/evardefine.ml +++ b/pretyping/evardefine.ml @@ -10,8 +10,9 @@ open Util open Pp open Names open Term -open Vars open Termops +open EConstr +open Vars open Namegen open Environ open Evd @@ -75,7 +76,6 @@ let idx = Namegen.default_dependent_ident let define_pure_evar_as_product evd evk = let open Context.Named.Declaration in - let open EConstr in let evi = Evd.find_undefined evd evk in let evenv = evar_env evi in let id = next_ident_away idx (ids_of_named_context (evar_context evi)) in @@ -105,19 +105,19 @@ let define_pure_evar_as_product evd evk = let evd3 = Evd.set_leq_sort evenv evd3 (Type prods) s in evd3, rng in - let prod = mkProd (Name id, EConstr.of_constr dom, EConstr.of_constr (subst_var id rng)) in + let rng = EConstr.of_constr rng in + let prod = mkProd (Name id, EConstr.of_constr dom, subst_var id rng) in let evd3 = Evd.define evk (EConstr.Unsafe.to_constr prod) evd2 in evd3,prod (* Refine an applied evar to a product and returns its instantiation *) let define_evar_as_product evd (evk,args) = - let open EConstr in let evd,prod = define_pure_evar_as_product evd evk in (* Quick way to compute the instantiation of evk with args *) let na,dom,rng = destProd evd prod in let evdom = mkEvar (fst (destEvar evd dom), args) in - let evrngargs = Array.cons (mkRel 1) (Array.map (Vars.lift 1) args) in + let evrngargs = Array.cons (mkRel 1) (Array.map (lift 1) args) in let evrng = mkEvar (fst (destEvar evd rng), evrngargs) in evd, mkProd (na, evdom, evrng) @@ -132,7 +132,6 @@ let define_evar_as_product evd (evk,args) = let define_pure_evar_as_lambda env evd evk = let open Context.Named.Declaration in - let open EConstr in let evi = Evd.find_undefined evd evk in let evenv = evar_env evi in let typ = EConstr.of_constr (Reductionops.whd_all evenv evd (EConstr.of_constr (evar_concl evi))) in @@ -147,23 +146,21 @@ let define_pure_evar_as_lambda env evd evk = let newenv = push_named (LocalAssum (id, dom)) evenv in let filter = Filter.extend 1 (evar_filter evi) in let src = evar_source evk evd1 in - let evd2,body = new_evar_unsafe newenv evd1 ~src (Vars.subst1 (mkVar id) rng) ~filter in - let lam = mkLambda (Name id, EConstr.of_constr dom, Vars.subst_var id (EConstr.of_constr body)) in + let evd2,body = new_evar_unsafe newenv evd1 ~src (subst1 (mkVar id) rng) ~filter in + let lam = mkLambda (Name id, EConstr.of_constr dom, subst_var id (EConstr.of_constr body)) in Evd.define evk (EConstr.Unsafe.to_constr lam) evd2, lam let define_evar_as_lambda env evd (evk,args) = - let open EConstr in let evd,lam = define_pure_evar_as_lambda env evd evk in (* Quick way to compute the instantiation of evk with args *) let na,dom,body = destLambda evd lam in - let evbodyargs = Array.cons (mkRel 1) (Array.map (Vars.lift 1) args) in + let evbodyargs = Array.cons (mkRel 1) (Array.map (lift 1) args) in let evbody = mkEvar (fst (destEvar evd body), evbodyargs) in evd, mkLambda (na, dom, evbody) let rec evar_absorb_arguments env evd (evk,args as ev) = function | [] -> evd,ev | a::l -> - let open EConstr in (* TODO: optimize and avoid introducing intermediate evars *) let evd,lam = define_pure_evar_as_lambda env evd evk in let _,_,body = destLambda evd lam in @@ -177,8 +174,9 @@ let define_evar_as_sort env evd (ev,args) = let evi = Evd.find_undefined evd ev in let s = Type u in let concl = Reductionops.whd_all (evar_env evi) evd (EConstr.of_constr evi.evar_concl) in - let sort = destSort concl in - let evd' = Evd.define ev (mkSort s) evd in + let concl = EConstr.of_constr concl in + let sort = destSort evd concl in + let evd' = Evd.define ev (Constr.mkSort s) evd in Evd.set_leq_sort env evd' (Type (Univ.super u)) sort, s (* Propagation of constraints through application and abstraction: @@ -187,7 +185,6 @@ let define_evar_as_sort env evd (ev,args) = an evar instantiate it with the product of 2 new evars. *) let split_tycon loc env evd tycon = - let open EConstr in let rec real_split evd c = let t = Reductionops.whd_all env evd c in match EConstr.kind evd (EConstr.of_constr t) with @@ -208,7 +205,7 @@ let split_tycon loc env evd tycon = evd', (n, mk_tycon dom, mk_tycon rng) let valcon_of_tycon x = x -let lift_tycon n = Option.map (EConstr.Vars.lift n) +let lift_tycon n = Option.map (lift n) let pr_tycon env = function None -> str "None" diff --git a/pretyping/evarsolve.ml b/pretyping/evarsolve.ml index b1fc7cbe9..b7db51d5c 100644 --- a/pretyping/evarsolve.ml +++ b/pretyping/evarsolve.ml @@ -6,14 +6,16 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +module CVars = Vars open Util open CErrors open Names open Term -open Vars open Environ open Termops open Evd +open EConstr +open Vars open Namegen open Retyping open Reductionops @@ -22,7 +24,6 @@ open Pretype_errors open Sigma.Notations let normalize_evar evd ev = - let open EConstr in match EConstr.kind evd (mkEvar ev) with | Evar (evk,args) -> (evk,args) | _ -> assert false @@ -50,7 +51,6 @@ let refresh_level evd s = let refresh_universes ?(status=univ_rigid) ?(onlyalg=false) ?(refreshset=false) pbty env evd t = - let open EConstr in let evdref = ref evd in let modified = ref false in let rec refresh status dir t = @@ -141,7 +141,6 @@ let add_conv_oriented_pb ?(tail=true) (pbty,env,t1,t2) evd = exception IllTypedInstance of env * EConstr.types * EConstr.types let recheck_applications conv_algo env evdref t = - let open EConstr in let rec aux env t = match EConstr.kind !evdref t with | App (f, args) -> @@ -154,7 +153,7 @@ let recheck_applications conv_algo env evdref t = | Prod (na, dom, codom) -> (match conv_algo env !evdref Reduction.CUMUL argsty.(i) dom with | Success evd -> evdref := evd; - aux (succ i) (Vars.subst1 args.(i) codom) + aux (succ i) (subst1 args.(i) codom) | UnifFailure (evd, reason) -> Pretype_errors.error_cannot_unify env evd ~reason (argsty.(i), dom)) | _ -> raise (IllTypedInstance (env, ty, argsty.(i))) @@ -221,7 +220,6 @@ let restrict_instance evd evk filter argsv = open Context.Rel.Declaration let noccur_evar env evd evk c = - let open EConstr in let cache = ref Int.Set.empty (* cache for let-ins *) in let rec occur_rec check_types (k, env as acc) c = match EConstr.kind evd c with @@ -234,10 +232,10 @@ let noccur_evar env evd evk c = if not (Int.Set.mem (i-k) !cache) then let decl = Environ.lookup_rel i env in if check_types then - (cache := Int.Set.add (i-k) !cache; occur_rec false acc (Vars.lift i (EConstr.of_constr (get_type decl)))); + (cache := Int.Set.add (i-k) !cache; occur_rec false acc (lift i (EConstr.of_constr (get_type decl)))); (match decl with | LocalAssum _ -> () - | LocalDef (_,b,_) -> cache := Int.Set.add (i-k) !cache; occur_rec false acc (Vars.lift i (EConstr.of_constr b))) + | LocalDef (_,b,_) -> cache := Int.Set.add (i-k) !cache; occur_rec false acc (lift i (EConstr.of_constr b))) | Proj (p,c) -> occur_rec true acc c | _ -> iter_with_full_binders evd (fun rd (k,env) -> (succ k, push_rel rd env)) (occur_rec check_types) acc c @@ -270,7 +268,6 @@ let compute_var_aliases sign sigma = sign Id.Map.empty let compute_rel_aliases var_aliases rels sigma = - let open EConstr in snd (List.fold_right (fun decl (n,aliases) -> (n-1, @@ -288,7 +285,7 @@ let compute_rel_aliases var_aliases rels sigma = try Int.Map.find (p+n) aliases with Not_found -> [] in Int.Map.add n (aliases_of_n@[mkRel (p+n)]) aliases | _ -> - Int.Map.add n [Vars.lift n (mkCast(t,DEFAULTcast,u))] aliases) + Int.Map.add n [lift n (mkCast(t,DEFAULTcast,u))] aliases) | LocalAssum _ -> aliases) ) rels @@ -301,10 +298,9 @@ let make_alias_map env sigma = (var_aliases,rel_aliases) let lift_aliases n (var_aliases,rel_aliases as aliases) = - let open EConstr in if Int.equal n 0 then aliases else (var_aliases, - Int.Map.fold (fun p l -> Int.Map.add (p+n) (List.map (Vars.lift n) l)) + Int.Map.fold (fun p l -> Int.Map.add (p+n) (List.map (lift n) l)) rel_aliases Int.Map.empty) let get_alias_chain_of sigma aliases x = match EConstr.kind sigma x with @@ -313,7 +309,6 @@ let get_alias_chain_of sigma aliases x = match EConstr.kind sigma x with | _ -> [] let normalize_alias_opt sigma aliases x = - let open EConstr in match get_alias_chain_of sigma aliases x with | [] -> None | a::_ when isRel sigma a || isVar sigma a -> Some a @@ -326,13 +321,11 @@ let normalize_alias sigma aliases x = | None -> x let normalize_alias_var sigma var_aliases id = - let open EConstr in destVar sigma (normalize_alias sigma (var_aliases,Int.Map.empty) (mkVar id)) let extend_alias sigma decl (var_aliases,rel_aliases) = - let open EConstr in let rel_aliases = - Int.Map.fold (fun n l -> Int.Map.add (n+1) (List.map (Vars.lift 1) l)) + Int.Map.fold (fun n l -> Int.Map.add (n+1) (List.map (lift 1) l)) rel_aliases Int.Map.empty in let rel_aliases = match decl with @@ -348,7 +341,7 @@ let extend_alias sigma decl (var_aliases,rel_aliases) = try Int.Map.find (p+1) rel_aliases with Not_found -> [] in Int.Map.add 1 (aliases_of_binder@[mkRel (p+1)]) rel_aliases | _ -> - Int.Map.add 1 [Vars.lift 1 t] rel_aliases) + Int.Map.add 1 [lift 1 t] rel_aliases) | LocalAssum _ -> rel_aliases in (var_aliases, rel_aliases) @@ -358,7 +351,6 @@ let expand_alias_once sigma aliases x = | l -> Some (List.last l) let expansions_of_var sigma aliases x = - let open EConstr in match get_alias_chain_of sigma aliases x with | [] -> [x] | a::_ as l when isRel sigma a || isVar sigma a -> x :: List.rev l @@ -379,7 +371,6 @@ let rec expand_vars_in_term_using sigma aliases t = match EConstr.kind sigma t w let expand_vars_in_term env sigma = expand_vars_in_term_using sigma (make_alias_map env sigma) let free_vars_and_rels_up_alias_expansion sigma aliases c = - let open EConstr in let acc1 = ref Int.Set.empty and acc2 = ref Id.Set.empty in let acc3 = ref Int.Set.empty and acc4 = ref Id.Set.empty in let cache_rel = ref Int.Set.empty and cache_var = ref Id.Set.empty in @@ -430,7 +421,7 @@ let rec expand_and_check_vars sigma aliases = function raise Exit module Constrhash = Hashtbl.Make - (struct type t = constr + (struct type t = Constr.constr let equal = Term.eq_constr let hash = hash_constr end) @@ -476,7 +467,6 @@ let remove_instance_local_defs evd evk args = (* Check if an applied evar "?X[args] l" is a Miller's pattern *) let find_unification_pattern_args env evd l t = - let open EConstr in if List.for_all (fun x -> isRel evd x || isVar evd x) l (* common failure case *) then let aliases = make_alias_map env evd in match (try Some (expand_and_check_vars evd aliases l) with Exit -> None) with @@ -488,7 +478,6 @@ let find_unification_pattern_args env evd l t = let is_unification_pattern_meta env evd nb m l t = (* Variables from context and rels > nb are implicitly all there *) (* so we need to be a rel <= nb *) - let open EConstr in if List.for_all (fun x -> isRel evd x && destRel evd x <= nb) l then match find_unification_pattern_args env evd l t with | Some _ as x when not (dependent evd (EConstr.mkMeta m) t) -> x @@ -497,7 +486,6 @@ let is_unification_pattern_meta env evd nb m l t = None let is_unification_pattern_evar env evd (evk,args) l t = - let open EConstr in if List.for_all (fun x -> isRel evd x || isVar evd x) l && noccur_evar env evd evk t then @@ -528,14 +516,13 @@ let is_unification_pattern (env,nb) evd f l t = dependency: ?X x = ?1 (?1 is a meta) will return \_.?1 while it should return \y. ?1{x\y} (non constant function if ?1 depends on x) (BB) *) let solve_pattern_eqn env sigma l c = - let open EConstr in let c' = List.fold_right (fun a c -> - let c' = subst_term sigma (Vars.lift 1 a) (Vars.lift 1 c) in + let c' = subst_term sigma (lift 1 a) (lift 1 c) in match EConstr.kind sigma a with (* Rem: if [a] links to a let-in, do as if it were an assumption *) | Rel n -> let open Context.Rel.Declaration in - let d = map_constr (lift n) (lookup_rel n env) in + let d = map_constr (CVars.lift n) (lookup_rel n env) in let c' = EConstr.of_constr c' in mkLambda_or_LetIn d c' | Var id -> @@ -604,7 +591,6 @@ let make_projectable_subst aliases sigma evi args = *) let define_evar_from_virtual_equation define_fun env evd src t_in_env ty_t_in_sign sign filter inst_in_env = - let open EConstr in let evd = Sigma.Unsafe.of_evar_map evd in let Sigma (evar_in_env, evd, _) = new_evar_instance sign evd (EConstr.of_constr ty_t_in_sign) ~filter ~src (List.map EConstr.Unsafe.to_constr inst_in_env) in let evd = Sigma.to_evar_map evd in @@ -637,7 +623,6 @@ let define_evar_from_virtual_equation define_fun env evd src t_in_env ty_t_in_si exception MorePreciseOccurCheckNeeeded let materialize_evar define_fun env evd k (evk1,args1) ty_in_env = - let open EConstr in if Evd.is_defined evd evk1 then (* Some circularity somewhere (see e.g. #3209) *) raise MorePreciseOccurCheckNeeeded; @@ -669,8 +654,8 @@ let materialize_evar define_fun env evd k (evk1,args1) ty_in_env = t_in_sign sign filter inst_in_env in evd, Context.Named.Declaration.LocalDef (id,b,t_in_sign) in (push_named_context_val d' sign, Filter.extend 1 filter, - (mkRel 1)::(List.map (Vars.lift 1) inst_in_env), - (mkRel 1)::(List.map (Vars.lift 1) inst_in_sign), + (mkRel 1)::(List.map (lift 1) inst_in_env), + (mkRel 1)::(List.map (lift 1) inst_in_sign), push_rel d env,evd,id::avoid)) rel_sign (sign1,filter1,Array.to_list args1,inst_in_sign,env1,evd,ids1) @@ -707,7 +692,7 @@ let find_projectable_constructor env evd cstr k args cstr_subst = List.filter (fun (args',id) -> (* is_conv is maybe too strong (and source of useless computation) *) (* (at least expansion of aliases is needed) *) - Array.for_all2 (fun c1 c2 -> is_conv env evd (EConstr.of_constr c1) (EConstr.of_constr c2)) args args') l in + Array.for_all2 (fun c1 c2 -> is_conv env evd c1 (EConstr.of_constr c2)) args args') l in List.map snd l with Not_found -> [] @@ -765,7 +750,6 @@ let rec assoc_up_to_alias sigma aliases y yc = function | _ -> if EConstr.eq_constr sigma yc c then id else raise Not_found let rec find_projectable_vars with_evars aliases sigma y subst = - let open EConstr in let yc = normalize_alias sigma aliases y in let is_projectable idc idcl subst' = (* First test if some [id] aliased to [idc] is bound to [y] in [subst] *) @@ -829,7 +813,6 @@ let rec find_solution_type evarenv = function let rec do_projection_effects define_fun env ty evd = function | ProjectVar -> evd | ProjectEvar ((evk,argsv),evi,id,p) -> - let open EConstr in let evd = Evd.define evk (Constr.mkVar id) evd in (* TODO: simplify constraints involving evk *) let evd = do_projection_effects define_fun env ty evd p in @@ -840,7 +823,7 @@ let rec do_projection_effects define_fun env ty evd = function one (however, regarding coercions, because t is obtained by unif, we know that no coercion can be inserted) *) let subst = make_pure_subst evi argsv in - let ty' = Vars.replace_vars subst (EConstr.of_constr evi.evar_concl) in + let ty' = replace_vars subst (EConstr.of_constr evi.evar_concl) in if isEvar evd ty' then define_fun env evd (Some false) (destEvar evd ty') ty else evd else evd @@ -875,14 +858,13 @@ type projectibility_status = let invert_arg_from_subst evd aliases k0 subst_in_env_extended_with_k_binders c_in_env_extended_with_k_binders = let effects = ref [] in - let open EConstr in let rec aux k t = match EConstr.kind evd t with | Rel i when i>k0+k -> aux' k (mkRel (i-k)) | Var id -> aux' k t | _ -> map_with_binders evd succ aux k t and aux' k t = - try EConstr.of_constr (project_with_effects aliases evd effects t subst_in_env_extended_with_k_binders) + try project_with_effects aliases evd effects t subst_in_env_extended_with_k_binders with Not_found -> match expand_alias_once evd aliases t with | None -> raise Not_found @@ -983,7 +965,8 @@ let closure_of_filter evd evk = function | LocalAssum _ -> false | LocalDef (_,c,_) -> - not (isRel c || isVar c) + let c = EConstr.of_constr c in + not (isRel evd c || isVar evd c) in let newfilter = Filter.map_along test filter (evar_context evi) in (* Now ensure that restriction is at least what is was originally *) @@ -1009,7 +992,6 @@ let restrict_hyps evd evk filter candidates = exception EvarSolvedWhileRestricting of evar_map * EConstr.constr let do_restrict_hyps evd (evk,args as ev) filter candidates = - let open EConstr in let filter,candidates = match filter with | None -> None,candidates | Some filter -> restrict_hyps evd evk filter candidates in @@ -1025,7 +1007,6 @@ let do_restrict_hyps evd (evk,args as ev) filter candidates = (* ?e is assumed to have no candidates *) let postpone_non_unique_projection env evd pbty (evk,argsv as ev) sols rhs = - let open EConstr in let rhs = expand_vars_in_term env evd rhs in let filter = restrict_upon_filter evd evk @@ -1162,7 +1143,6 @@ exception EvarSolvedOnTheFly of evar_map * EConstr.constr (* Try to project evk1[argsv1] on evk2[argsv2], if [ev1] is a pattern on the common domain of definition *) let project_evar_on_evar force g env evd aliases k2 pbty (evk1,argsv1 as ev1) (evk2,argsv2 as ev2) = - let open EConstr in (* Apply filtering on ev1 so that fvs(ev1) are in fvs(ev2). *) let fvs2 = free_vars_and_rels_up_alias_expansion evd aliases (mkEvar ev2) in let filter1 = restrict_upon_filter evd evk1 @@ -1210,7 +1190,6 @@ let update_evar_source ev1 ev2 evd = let solve_evar_evar_l2r force f g env evd aliases pbty ev1 (evk2,_ as ev2) = try - let open EConstr in let evd,body = project_evar_on_evar force g env evd aliases 0 pbty ev1 ev2 in let evd' = Evd.define evk2 (EConstr.Unsafe.to_constr body) evd in let evd' = update_evar_source (fst (destEvar evd body)) evk2 evd' in @@ -1230,7 +1209,6 @@ let preferred_orientation evd evk1 evk2 = | _ -> true let solve_evar_evar_aux force f g env evd pbty (evk1,args1 as ev1) (evk2,args2 as ev2) = - let open EConstr in let aliases = make_alias_map env evd in if preferred_orientation evd evk1 evk2 then try solve_evar_evar_l2r force f g env evd aliases (opp_problem pbty) ev2 ev1 @@ -1248,6 +1226,7 @@ let solve_evar_evar_aux force f g env evd pbty (evk1,args1 as ev1) (evk2,args2 a let solve_evar_evar ?(force=false) f g env evd pbty (evk1,args1 as ev1) (evk2,args2 as ev2) = let pbty = if force then None else pbty in let evi = Evd.find evd evk1 in + let downcast evk t evd = downcast evk (EConstr.Unsafe.to_constr t) evd in let evd = try (* ?X : Π Δ. Type i = ?Y : Π Δ'. Type j. @@ -1289,7 +1268,6 @@ type conv_fun_bool = * depend on these args). *) let solve_refl ?(can_drop=false) conv_algo env evd pbty evk argsv1 argsv2 = - let open EConstr in let evdref = ref evd in if Array.equal (fun c1 c2 -> e_eq_constr_univs evdref (EConstr.Unsafe.to_constr c1) (EConstr.Unsafe.to_constr c2) ) argsv1 argsv2 then !evdref else (* Filter and restrict if needed *) @@ -1350,7 +1328,7 @@ let occur_evar_upto_types sigma n c = else ( seen := Evar.Set.add sp !seen; Option.iter occur_rec (existential_opt_value sigma e); - occur_rec (existential_type sigma e)) + occur_rec (Evd.existential_type sigma e)) | _ -> Constr.iter occur_rec c in try occur_rec c; false with Occur -> true @@ -1385,7 +1363,6 @@ exception OccurCheckIn of evar_map * EConstr.constr exception MetaOccurInBodyInternal let rec invert_definition conv_algo choose env evd pbty (evk,argsv as ev) rhs = - let open EConstr in let aliases = make_alias_map env evd in let evdref = ref evd in let progress = ref false in @@ -1441,7 +1418,7 @@ let rec invert_definition conv_algo choose env evd pbty (evk,argsv as ev) rhs = | LocalAssum _ -> project_variable (mkRel (i-k)) | LocalDef (_,b,_) -> try project_variable (mkRel (i-k)) - with NotInvertibleUsingOurAlgorithm _ -> imitate envk (Vars.lift i (EConstr.of_constr b))) + with NotInvertibleUsingOurAlgorithm _ -> imitate envk (lift i (EConstr.of_constr b))) | Var id -> (match Environ.lookup_named id env' with | LocalAssum _ -> project_variable t @@ -1449,13 +1426,13 @@ let rec invert_definition conv_algo choose env evd pbty (evk,argsv as ev) rhs = try project_variable t with NotInvertibleUsingOurAlgorithm _ -> imitate envk (EConstr.of_constr b)) | LetIn (na,b,u,c) -> - imitate envk (Vars.subst1 b c) + imitate envk (subst1 b c) | Evar (evk',args' as ev') -> if Evar.equal evk evk' then raise (OccurCheckIn (evd,rhs)); (* Evar/Evar problem (but left evar is virtual) *) let aliases = lift_aliases k aliases in (try - let ev = (evk,Array.map (Vars.lift k) argsv) in + let ev = (evk,Array.map (lift k) argsv) in let evd,body = project_evar_on_evar false conv_algo env' !evdref aliases k None ev' ev in evdref := evd; body @@ -1487,8 +1464,9 @@ let rec invert_definition conv_algo choose env evd pbty (evk,argsv as ev) rhs = progress := true; match let c,args = decompose_app_vect !evdref t in + let args = Array.map EConstr.of_constr args in match EConstr.kind !evdref (EConstr.of_constr c) with - | Construct (cstr,u) when Vars.noccur_between !evdref 1 k t -> + | Construct (cstr,u) when noccur_between !evdref 1 k t -> (* This is common case when inferring the return clause of match *) (* (currently rudimentary: we do not treat the case of multiple *) (* possible inversions; we do not treat overlap with a possible *) @@ -1533,7 +1511,7 @@ let rec invert_definition conv_algo choose env evd pbty (evk,argsv as ev) rhs = | _ -> false in is_id_subst filter_ctxt (Array.to_list argsv) && - Vars.closed0 evd rhs && + closed0 evd rhs && Idset.subset (collect_vars evd rhs) !names in let body = @@ -1659,7 +1637,6 @@ let reconsider_conv_pbs conv_algo evd = (* Rq: uncomplete algorithm if pbty = CONV_X_LEQ ! *) let solve_simple_eqn conv_algo ?(choose=false) env evd (pbty,(evk1,args1 as ev1),t2) = - let open EConstr in try let t2 = EConstr.of_constr (whd_betaiota evd t2) in (* includes whd_evar *) let evd = evar_define conv_algo ~choose env evd pbty ev1 t2 in diff --git a/pretyping/patternops.ml b/pretyping/patternops.ml index d473f41bd..ffd6e73fa 100644 --- a/pretyping/patternops.ml +++ b/pretyping/patternops.ml @@ -163,7 +163,7 @@ let pattern_of_constr env sigma t = | Ind (sp,u) -> PRef (canonical_gr (IndRef sp)) | Construct (sp,u) -> PRef (canonical_gr (ConstructRef sp)) | Proj (p, c) -> - pattern_of_constr env (EConstr.of_constr (Retyping.expand_projection env sigma p c [])) + pattern_of_constr env (Retyping.expand_projection env sigma p c []) | Evar (evk,ctxt as ev) -> (match snd (Evd.evar_source evk sigma) with | Evar_kinds.MatchingVar (b,id) -> diff --git a/pretyping/reductionops.ml b/pretyping/reductionops.ml index 510417879..0e0b80744 100644 --- a/pretyping/reductionops.ml +++ b/pretyping/reductionops.ml @@ -10,10 +10,11 @@ open CErrors open Util open Names open Term -open Vars open Termops open Univ open Evd +open EConstr +open Vars open Environ open Context.Rel.Declaration @@ -574,7 +575,7 @@ struct zip (best_state sigma (constr_of_cst_member cst (params @ (append_app [|f|] s))) cst_l) | f, (Cst (cst,_,_,params,_)::s) -> zip (constr_of_cst_member cst (params @ (append_app [|f|] s))) - | f, (Shift n::s) -> zip (Vars.lift n f, s) + | f, (Shift n::s) -> zip (lift n f, s) | f, (Proj (n,m,p,cst_l)::s) when refold -> zip (best_state sigma (mkProj (p,f),s) cst_l) | f, (Proj (n,m,p,_)::s) -> zip (mkProj (p,f),s) @@ -585,18 +586,18 @@ struct end (** The type of (machine) states (= lambda-bar-calculus' cuts) *) -type state = EConstr.t * EConstr.t Stack.t +type state = constr * constr Stack.t -type contextual_reduction_function = env -> evar_map -> EConstr.t -> constr +type contextual_reduction_function = env -> evar_map -> constr -> Constr.constr type reduction_function = contextual_reduction_function -type local_reduction_function = evar_map -> EConstr.t -> constr -type e_reduction_function = { e_redfun : 'r. env -> 'r Sigma.t -> EConstr.t -> (constr, 'r) Sigma.sigma } +type local_reduction_function = evar_map -> constr -> Constr.constr +type e_reduction_function = { e_redfun : 'r. env -> 'r Sigma.t -> constr -> (Constr.constr, 'r) Sigma.sigma } type contextual_stack_reduction_function = - env -> evar_map -> EConstr.t -> EConstr.t * EConstr.t list + env -> evar_map -> constr -> constr * constr list type stack_reduction_function = contextual_stack_reduction_function type local_stack_reduction_function = - evar_map -> EConstr.t -> EConstr.t * EConstr.t list + evar_map -> constr -> constr * constr list type contextual_state_reduction_function = env -> evar_map -> state -> state @@ -639,7 +640,7 @@ let local_strong whdfun sigma = let rec strong_prodspine redfun sigma c = let x = EConstr.of_constr (redfun sigma c) in match EConstr.kind sigma x with - | Prod (na,a,b) -> mkProd (na, EConstr.Unsafe.to_constr a,strong_prodspine redfun sigma b) + | Prod (na,a,b) -> EConstr.Unsafe.to_constr (mkProd (na,a,EConstr.of_constr (strong_prodspine redfun sigma b))) | _ -> EConstr.Unsafe.to_constr x (*************************************) @@ -656,7 +657,7 @@ let apply_subst recfun env sigma refold cst_l t stack = | Some (h,stacktl), Lambda (_,_,c) -> let cst_l' = if refold then Cst_stack.add_param h cst_l else cst_l in aux (h::env) cst_l' c stacktl - | _ -> recfun sigma cst_l (EConstr.Vars.substl env t, stack) + | _ -> recfun sigma cst_l (substl env t, stack) in aux env cst_l t stack let stacklam recfun env sigma t stack = @@ -673,8 +674,8 @@ let beta_applist sigma (c,l) = (* Iota reduction tools *) type 'a miota_args = { - mP : EConstr.t; (* the result type *) - mconstr : EConstr.t; (* the constructor *) + mP : constr; (* the result type *) + mconstr : constr; (* the constructor *) mci : case_info; (* special info to re-build pattern *) mcargs : 'a list; (* the constructor's arguments *) mlf : 'a array } (* the branch code vector *) @@ -696,7 +697,6 @@ let reducible_mind_case sigma c = match EConstr.kind sigma c with let magicaly_constant_of_fixbody env sigma reference bd = function | Name.Anonymous -> bd | Name.Name id -> - let open EConstr in try let (cst_mod,cst_sect,_) = Constant.repr3 reference in let cst = Constant.make3 cst_mod cst_sect (Label.of_id id) in @@ -719,7 +719,6 @@ let magicaly_constant_of_fixbody env sigma reference bd = function | Not_found -> bd let contract_cofix ?env sigma ?reference (bodynum,(names,types,bodies as typedbodies)) = - let open EConstr in let nbodies = Array.length bodies in let make_Fi j = let ind = nbodies-j-1 in @@ -733,11 +732,10 @@ let contract_cofix ?env sigma ?reference (bodynum,(names,types,bodies as typedbo | None -> bd | Some r -> magicaly_constant_of_fixbody e sigma r bd names.(ind) in let closure = List.init nbodies make_Fi in - Vars.substl closure bodies.(bodynum) + substl closure bodies.(bodynum) (** Similar to the "fix" case below *) let reduce_and_refold_cofix recfun env sigma refold cst_l cofix sk = - let open EConstr in let raw_answer = let env = if refold then Some env else None in contract_cofix ?env sigma ?reference:(Cst_stack.reference sigma cst_l) cofix in @@ -749,7 +747,6 @@ let reduce_and_refold_cofix recfun env sigma refold cst_l cofix sk = [] sigma refold Cst_stack.empty raw_answer sk let reduce_mind_case sigma mia = - let open EConstr in match EConstr.kind sigma mia.mconstr with | Construct ((ind_sp,i),u) -> (* let ncargs = (fst mia.mci).(i-1) in*) @@ -764,7 +761,6 @@ let reduce_mind_case sigma mia = Bi[Fj --> FIX[nl;j](A1...Ak;[F1...Fk]{B1...Bk})] *) let contract_fix ?env sigma ?reference ((recindices,bodynum),(names,types,bodies as typedbodies)) = - let open EConstr in let nbodies = Array.length recindices in let make_Fi j = let ind = nbodies-j-1 in @@ -778,14 +774,13 @@ let contract_fix ?env sigma ?reference ((recindices,bodynum),(names,types,bodies | None -> bd | Some r -> magicaly_constant_of_fixbody e sigma r bd names.(ind) in let closure = List.init nbodies make_Fi in - Vars.substl closure bodies.(bodynum) + substl closure bodies.(bodynum) (** First we substitute the Rel bodynum by the fixpoint and then we try to replace the fixpoint by the best constant from [cst_l] Other rels are directly substituted by constants "magically found from the context" in contract_fix *) let reduce_and_refold_fix recfun env sigma refold cst_l fix sk = - let open EConstr in let raw_answer = let env = if refold then None else Some env in contract_fix ?env sigma ?reference:(Cst_stack.reference sigma cst_l) fix in @@ -830,7 +825,6 @@ let _ = Goptions.declare_bool_option { } let equal_stacks sigma (x, l) (y, l') = - let open EConstr in let f_equal (x,lft1) (y,lft2) = eq_constr sigma (Vars.lift lft1 x) (Vars.lift lft2 y) in let eq_fix (a,b) (c,d) = f_equal (mkFix a, b) (mkFix c, d) in match Stack.equal f_equal eq_fix l l' with @@ -838,7 +832,6 @@ let equal_stacks sigma (x, l) (y, l') = | Some (lft1,lft2) -> f_equal (x, lft1) (y, lft2) let rec whd_state_gen ?csts ~refold ~tactic_mode flags env sigma = - let open EConstr in let open Context.Named.Declaration in let rec whrec cst_l (x, stack) = let () = if !debug_RAKAM then @@ -859,7 +852,7 @@ let rec whd_state_gen ?csts ~refold ~tactic_mode flags env sigma = match c0 with | Rel n when CClosure.RedFlags.red_set flags CClosure.RedFlags.fDELTA -> (match lookup_rel n env with - | LocalDef (_,body,_) -> whrec Cst_stack.empty (EConstr.of_constr (lift n body), stack) + | LocalDef (_,body,_) -> whrec Cst_stack.empty (lift n (EConstr.of_constr body), stack) | _ -> fold ()) | Var id when CClosure.RedFlags.red_set flags (CClosure.RedFlags.fVAR id) -> (match lookup_named id env with @@ -977,7 +970,7 @@ let rec whd_state_gen ?csts ~refold ~tactic_mode flags env sigma = | Rel 1, [] -> let lc = Array.sub cl 0 (napp-1) in let u = if Int.equal napp 1 then f else mkApp (f,lc) in - if Vars.noccurn sigma 1 u then (EConstr.of_constr (pop u),Stack.empty),Cst_stack.empty else fold () + if noccurn sigma 1 u then (EConstr.of_constr (pop u),Stack.empty),Cst_stack.empty else fold () | _ -> fold () else fold () | _ -> fold ()) @@ -1054,7 +1047,6 @@ let rec whd_state_gen ?csts ~refold ~tactic_mode flags env sigma = (** reduction machine without global env and refold machinery *) let local_whd_state_gen flags sigma = - let open EConstr in let rec whrec (x, stack) = let c0 = EConstr.kind sigma x in let s = (EConstr.of_kind c0, stack) in @@ -1077,7 +1069,7 @@ let local_whd_state_gen flags sigma = | Rel 1, [] -> let lc = Array.sub cl 0 (napp-1) in let u = if Int.equal napp 1 then f else mkApp (f,lc) in - if Vars.noccurn sigma 1 u then (EConstr.of_constr (pop u),Stack.empty) else s + if noccurn sigma 1 u then (EConstr.of_constr (pop u),Stack.empty) else s | _ -> s else s | _ -> s) @@ -1276,7 +1268,7 @@ let f_conv_leq ?l2r ?reds env ?evars x y = let inj = EConstr.Unsafe.to_constr in Reduction.conv_leq ?l2r ?reds env ?evars (inj x) (inj y) -let test_trans_conversion (f: EConstr.t Reduction.extended_conversion_function) reds env sigma x y = +let test_trans_conversion (f: constr Reduction.extended_conversion_function) reds env sigma x y = try let evars ev = safe_evar_value sigma ev in let _ = f ~reds env ~evars:(evars, Evd.universes sigma) x y in @@ -1368,9 +1360,8 @@ let default_plain_instance_ident = Id.of_string "H" (* Try to replace all metas. Does not replace metas in the metas' values * Differs from (strong whd_meta). *) let plain_instance sigma s c = - let open EConstr in let rec irec n u = match EConstr.kind sigma u with - | Meta p -> (try Vars.lift n (Metamap.find p s) with Not_found -> u) + | Meta p -> (try lift n (Metamap.find p s) with Not_found -> u) | App (f,l) when isCast sigma f -> let (f,_,t) = destCast sigma f in let l' = CArray.Fun1.smartmap irec n l in @@ -1382,13 +1373,13 @@ let plain_instance sigma s c = (try let g = Metamap.find p s in match EConstr.kind sigma g with | App _ -> - let l' = CArray.Fun1.smartmap Vars.lift 1 l' in + let l' = CArray.Fun1.smartmap lift 1 l' in mkLetIn (Name default_plain_instance_ident,g,t,mkApp(mkRel 1, l')) | _ -> mkApp (g,l') with Not_found -> mkApp (f,l')) | _ -> mkApp (irec n f,l')) | Cast (m,_,_) when isMeta sigma m -> - (try Vars.lift n (Metamap.find (destMeta sigma m) s) with Not_found -> u) + (try lift n (Metamap.find (destMeta sigma m) s) with Not_found -> u) | _ -> map_with_binders sigma succ irec n u in @@ -1440,9 +1431,8 @@ let instance sigma s c = * error message. *) let hnf_prod_app env sigma t n = - let open EConstr in match EConstr.kind sigma (EConstr.of_constr (whd_all env sigma t)) with - | Prod (_,_,b) -> EConstr.Unsafe.to_constr (Vars.subst1 n b) + | Prod (_,_,b) -> EConstr.Unsafe.to_constr (subst1 n b) | _ -> anomaly ~label:"hnf_prod_app" (Pp.str "Need a product") let hnf_prod_appvect env sigma t nl = @@ -1452,9 +1442,8 @@ let hnf_prod_applist env sigma t nl = List.fold_left (fun acc t -> hnf_prod_app env sigma (EConstr.of_constr acc) t) (EConstr.Unsafe.to_constr t) nl let hnf_lam_app env sigma t n = - let open EConstr in match EConstr.kind sigma (EConstr.of_constr (whd_all env sigma t)) with - | Lambda (_,_,b) -> EConstr.Unsafe.to_constr (Vars.subst1 n b) + | Lambda (_,_,b) -> EConstr.Unsafe.to_constr (subst1 n b) | _ -> anomaly ~label:"hnf_lam_app" (Pp.str "Need an abstraction") let hnf_lam_appvect env sigma t nl = @@ -1544,7 +1533,6 @@ let is_sort env sigma t = of case/fix (heuristic used by evar_conv) *) let whd_betaiota_deltazeta_for_iota_state ts env sigma csts s = - let open EConstr in let refold = get_refolding_in_reduction () in let tactic_mode = false in let rec whrec csts s = @@ -1586,7 +1574,6 @@ let is_arity env sigma c = (* Metas *) let meta_value evd mv = - let open EConstr in let rec valrec mv = match meta_opt_fvalue evd mv with | Some (b,_) -> @@ -1617,54 +1604,58 @@ let meta_reducible_instance evd b = in let metas = Metaset.fold fold fm Metamap.empty in let rec irec u = - let u = whd_betaiota Evd.empty (EConstr.of_constr u) (** FIXME *) in - match kind_of_term u with - | Case (ci,p,c,bl) when EConstr.isMeta evd (EConstr.of_constr (strip_outer_cast evd (EConstr.of_constr c))) -> - let m = destMeta (strip_outer_cast evd (EConstr.of_constr c)) in + let u = whd_betaiota Evd.empty u (** FIXME *) in + let u = EConstr.of_constr u in + match EConstr.kind evd u with + | Case (ci,p,c,bl) when EConstr.isMeta evd (EConstr.of_constr (strip_outer_cast evd c)) -> + let m = destMeta evd (EConstr.of_constr (strip_outer_cast evd c)) in (match try let g, s = Metamap.find m metas in + let g = EConstr.of_constr g in let is_coerce = match s with CoerceToType -> true | _ -> false in - if isConstruct g || not is_coerce then Some g else None + if isConstruct evd g || not is_coerce then Some g else None with Not_found -> None with | Some g -> irec (mkCase (ci,p,g,bl)) | None -> mkCase (ci,irec p,c,Array.map irec bl)) - | App (f,l) when EConstr.isMeta evd (EConstr.of_constr (strip_outer_cast evd (EConstr.of_constr f))) -> - let m = destMeta (strip_outer_cast evd (EConstr.of_constr f)) in + | App (f,l) when EConstr.isMeta evd (EConstr.of_constr (strip_outer_cast evd f)) -> + let m = destMeta evd (EConstr.of_constr (strip_outer_cast evd f)) in (match try let g, s = Metamap.find m metas in + let g = EConstr.of_constr g in let is_coerce = match s with CoerceToType -> true | _ -> false in - if isLambda g || not is_coerce then Some g else None + if isLambda evd g || not is_coerce then Some g else None with Not_found -> None with | Some g -> irec (mkApp (g,l)) | None -> mkApp (f,Array.map irec l)) | Meta m -> (try let g, s = Metamap.find m metas in + let g = EConstr.of_constr g in let is_coerce = match s with CoerceToType -> true | _ -> false in if not is_coerce then irec g else u with Not_found -> u) - | Proj (p,c) when isMeta c || isCast c && isMeta (pi1 (destCast c)) -> - let m = try destMeta c with _ -> destMeta (pi1 (destCast c)) in + | Proj (p,c) when isMeta evd c || isCast evd c && isMeta evd (pi1 (destCast evd c)) -> + let m = try destMeta evd c with _ -> destMeta evd (pi1 (destCast evd c)) in (match try let g, s = Metamap.find m metas in + let g = EConstr.of_constr g in let is_coerce = match s with CoerceToType -> true | _ -> false in - if isConstruct g || not is_coerce then Some g else None + if isConstruct evd g || not is_coerce then Some g else None with Not_found -> None with | Some g -> irec (mkProj (p,g)) | None -> mkProj (p,c)) - | _ -> Constr.map irec u + | _ -> EConstr.map evd irec u in if Metaset.is_empty fm then (* nf_betaiota? *) b.rebus - else irec b.rebus + else EConstr.Unsafe.to_constr (irec (EConstr.of_constr b.rebus)) let head_unfold_under_prod ts env sigma c = - let open EConstr in let unfold (cst,u as cstu) = if Cpred.mem cst (snd ts) then match constant_opt_value_in env cstu with @@ -1682,12 +1673,11 @@ let head_unfold_under_prod ts env sigma c = EConstr.Unsafe.to_constr (aux c) let betazetaevar_applist sigma n c l = - let open EConstr in let rec stacklam n env t stack = - if Int.equal n 0 then applist (Vars.substl env t, stack) else + if Int.equal n 0 then applist (substl env t, stack) else match EConstr.kind sigma t, stack with | Lambda(_,_,c), arg::stacktl -> stacklam (n-1) (arg::env) c stacktl - | LetIn(_,b,_,c), _ -> stacklam (n-1) (Vars.substl env b::env) c stack - | Evar _, _ -> applist (Vars.substl env t, stack) + | LetIn(_,b,_,c), _ -> stacklam (n-1) (substl env b::env) c stack + | Evar _, _ -> applist (substl env t, stack) | _ -> anomaly (Pp.str "Not enough lambda/let's") in EConstr.Unsafe.to_constr (stacklam n [] c l) diff --git a/pretyping/retyping.ml b/pretyping/retyping.ml index a7ccf98a6..6f03fc462 100644 --- a/pretyping/retyping.ml +++ b/pretyping/retyping.ml @@ -10,13 +10,14 @@ open Pp open CErrors open Util open Term -open Vars open Inductive open Inductiveops open Names open Reductionops open Environ open Termops +open EConstr +open Vars open Arguments_renaming open Context.Rel.Declaration @@ -58,7 +59,6 @@ let local_def (na, b, t) = LocalDef (na, inj b, inj t) let get_type_from_constraints env sigma t = - let open EConstr in if isEvar sigma (EConstr.of_constr (fst (decompose_app_vect sigma t))) then match List.map_filter (fun (pbty,env,t1,t2) -> @@ -74,19 +74,17 @@ let get_type_from_constraints env sigma t = let rec subst_type env sigma typ = function | [] -> typ | h::rest -> - let open EConstr in match EConstr.kind sigma (EConstr.of_constr (whd_all env sigma typ)) with - | Prod (na,c1,c2) -> subst_type env sigma (Vars.subst1 h c2) rest + | Prod (na,c1,c2) -> subst_type env sigma (subst1 h c2) rest | _ -> retype_error NonFunctionalConstruction (* If ft is the type of f which itself is applied to args, *) (* [sort_of_atomic_type] computes ft[args] which has to be a sort *) let sort_of_atomic_type env sigma ft args = - let open EConstr in let rec concl_of_arity env n ar args = match EConstr.kind sigma (EConstr.of_constr (whd_all env sigma ar)), args with - | Prod (na, t, b), h::l -> concl_of_arity (push_rel (local_def (na, Vars.lift n h, t)) env) (n + 1) b l + | Prod (na, t, b), h::l -> concl_of_arity (push_rel (local_def (na, lift n h, t)) env) (n + 1) b l | Sort s, [] -> s | _ -> retype_error NotASort in concl_of_arity env 0 ft (Array.to_list args) @@ -101,7 +99,6 @@ let decomp_sort env sigma t = | _ -> retype_error NotASort let retype ?(polyprop=true) sigma = - let open EConstr in let rec type_of env cstr = match EConstr.kind sigma cstr with | Meta n -> @@ -109,7 +106,7 @@ let retype ?(polyprop=true) sigma = with Not_found -> retype_error (BadMeta n)) | Rel n -> let ty = EConstr.of_constr (RelDecl.get_type (lookup_rel n env)) in - Vars.lift n ty + lift n ty | Var id -> type_of_var env id | Const cst -> EConstr.of_constr (rename_type_of_constant env cst) | Evar (evk, args) -> EConstr.of_constr (Evd.existential_type sigma (evk, Array.map EConstr.Unsafe.to_constr args)) @@ -133,7 +130,7 @@ let retype ?(polyprop=true) sigma = | Lambda (name,c1,c2) -> mkProd (name, c1, type_of (push_rel (local_assum (name,c1)) env) c2) | LetIn (name,b,c1,c2) -> - Vars.subst1 b (type_of (push_rel (local_def (name,b,c1)) env) c2) + subst1 b (type_of (push_rel (local_def (name,b,c1)) env) c2) | Fix ((_,i),(_,tys,_)) -> tys.(i) | CoFix (i,(_,tys,_)) -> tys.(i) | App(f,args) when is_template_polymorphic env sigma f -> @@ -257,11 +254,11 @@ let sorts_of_context env evc ctxt = snd (aux ctxt) let expand_projection env sigma pr c args = - let inj = EConstr.Unsafe.to_constr in let ty = get_type_of ~lax:true env sigma c in let (i,u), ind_args = try Inductiveops.find_mrectype env sigma (EConstr.of_constr ty) with Not_found -> retype_error BadRecursiveType in + let ind_args = List.map EConstr.of_constr ind_args in mkApp (mkConstU (Projection.constant pr,u), - Array.of_list (ind_args @ (inj c :: List.map inj args))) + Array.of_list (ind_args @ (c :: args))) diff --git a/pretyping/retyping.mli b/pretyping/retyping.mli index c84403890..a20b11b76 100644 --- a/pretyping/retyping.mli +++ b/pretyping/retyping.mli @@ -45,6 +45,6 @@ val type_of_global_reference_knowing_conclusion : val sorts_of_context : env -> evar_map -> Context.Rel.t -> sorts list -val expand_projection : env -> evar_map -> Names.projection -> EConstr.constr -> EConstr.constr list -> constr +val expand_projection : env -> evar_map -> Names.projection -> EConstr.constr -> EConstr.constr list -> EConstr.constr val print_retype_error : retype_error -> Pp.std_ppcmds diff --git a/pretyping/tacred.ml b/pretyping/tacred.ml index 5b8eaa50b..ae53c12ae 100644 --- a/pretyping/tacred.ml +++ b/pretyping/tacred.ml @@ -11,10 +11,11 @@ open CErrors open Util open Names open Term -open Vars open Libnames open Globnames open Termops +open EConstr +open Vars open Find_subterm open Namegen open Environ @@ -88,7 +89,6 @@ let evaluable_reference_eq sigma r1 r2 = match r1, r2 with | _ -> false let mkEvalRef ref u = - let open EConstr in match ref with | EvalConst cst -> mkConstU (cst,u) | EvalVar id -> mkVar id @@ -109,7 +109,6 @@ let destEvalRefU sigma c = match EConstr.kind sigma c with | _ -> anomaly (Pp.str "Not an unfoldable reference") let unsafe_reference_opt_value env sigma eval = - let open EConstr in match eval with | EvalConst cst -> (match (lookup_constant cst env).Declarations.const_body with @@ -118,20 +117,19 @@ let unsafe_reference_opt_value env sigma eval = | EvalVar id -> env |> lookup_named id |> NamedDecl.get_value |> Option.map EConstr.of_constr | EvalRel n -> - env |> lookup_rel n |> RelDecl.map_value (lift n) |> RelDecl.get_value |> Option.map EConstr.of_constr + env |> lookup_rel n |> RelDecl.get_value |> Option.map (EConstr.of_constr %> lift n) | EvalEvar ev -> match EConstr.kind sigma (mkEvar ev) with | Evar _ -> None | c -> Some (EConstr.of_kind c) let reference_opt_value env sigma eval u = - let open EConstr in match eval with | EvalConst cst -> Option.map EConstr.of_constr (constant_opt_value_in env (cst,u)) | EvalVar id -> env |> lookup_named id |> NamedDecl.get_value |> Option.map EConstr.of_constr | EvalRel n -> - env |> lookup_rel n |> RelDecl.map_value (lift n) |> RelDecl.get_value |> Option.map EConstr.of_constr + env |> lookup_rel n |> RelDecl.get_value |> Option.map (EConstr.of_constr %> lift n) | EvalEvar ev -> match EConstr.kind sigma (mkEvar ev) with | Evar _ -> None @@ -187,7 +185,6 @@ let eval_table = Summary.ref (Cmap.empty : frozen) ~name:"evaluation" *) let check_fix_reversibility sigma labs args ((lv,i),(_,tys,bds)) = - let open EConstr in let n = List.length labs in let nargs = List.length args in if nargs > n then raise Elimconst; @@ -232,7 +229,6 @@ let check_fix_reversibility sigma labs args ((lv,i),(_,tys,bds)) = let invert_name labs l na0 env sigma ref = function | Name id -> - let open EConstr in let minfxargs = List.length l in begin match na0 with | Name id' when Id.equal id' id -> @@ -276,7 +272,6 @@ let local_def (na, b, t) = LocalDef (na, inj b, inj t) let compute_consteval_direct env sigma ref = - let open EConstr in let rec srec env n labs onlyproj c = let c',l = whd_betadeltazeta_stack env sigma c in match EConstr.kind sigma c' with @@ -295,7 +290,6 @@ let compute_consteval_direct env sigma ref = | Some c -> srec env 0 [] false c let compute_consteval_mutual_fix env sigma ref = - let open EConstr in let rec srec env minarg labs ref c = let c',l = whd_betalet_stack sigma c in let nargs = List.length l in @@ -367,7 +361,6 @@ let reference_eval env sigma = function let x = Name default_dependent_ident let make_elim_fun (names,(nbfix,lv,n)) u largs = - let open EConstr in let lu = List.firstn n largs in let p = List.length lv in let lyi = List.map fst lv in @@ -393,7 +386,7 @@ let make_elim_fun (names,(nbfix,lv,n)) u largs = (* [f] is convertible to [Fix(recindices,bodynum),bodyvect)]: do so that the reduction uses this extra information *) -let dummy = mkProp +let dummy = Constr.mkProp let vfx = Id.of_string "_expanded_fix_" let vfun = Id.of_string "_eliminator_function_" let venv = let open Context.Named.Declaration in @@ -403,7 +396,6 @@ let venv = let open Context.Named.Declaration in as a problem variable: an evar that can be instantiated either by vfx (expanded fixpoint) or vfun (named function). *) let substl_with_function subst sigma constr = - let open EConstr in let evd = ref sigma in let minargs = ref Evar.Map.empty in let v = Array.of_list subst in @@ -431,8 +423,7 @@ exception Partial reduction is solved by the expanded fix term. *) let solve_arity_problem env sigma fxminargs c = let evm = ref sigma in - let set_fix i = evm := Evd.define i (mkVar vfx) !evm in - let open EConstr in + let set_fix i = evm := Evd.define i (Constr.mkVar vfx) !evm in let rec check strict c = let c' = EConstr.of_constr (whd_betaiotazeta sigma c) in let (h,rcargs) = decompose_app_vect sigma c' in @@ -491,7 +482,6 @@ let reduce_fix whdfun sigma fix stack = let contract_fix_use_function env sigma f ((recindices,bodynum),(_names,_types,bodies as typedbodies)) = - let open EConstr in let nbodies = Array.length recindices in let make_Fi j = (mkFix((recindices,j),typedbodies), f j) in let lbodies = List.init nbodies make_Fi in @@ -515,7 +505,6 @@ let reduce_fix_use_function env sigma f whfun fix stack = let contract_cofix_use_function env sigma f (bodynum,(_names,_,bodies as typedbodies)) = - let open EConstr in let nbodies = Array.length bodies in let make_Fi j = (mkCoFix(j,typedbodies), f j) in let subbodies = List.init nbodies make_Fi in @@ -523,7 +512,6 @@ let contract_cofix_use_function env sigma f sigma (EConstr.of_constr (nf_beta sigma bodies.(bodynum))) let reduce_mind_case_use_function func env sigma mia = - let open EConstr in match EConstr.kind sigma mia.mconstr with | Construct ((ind_sp,i),u) -> let real_cargs = List.skipn mia.mci.ci_npar mia.mcargs in @@ -573,11 +561,10 @@ let match_eval_ref_value env sigma constr = | Var id when is_evaluable env (EvalVarRef id) -> env |> lookup_named id |> NamedDecl.get_value |> Option.map EConstr.of_constr | Rel n -> - env |> lookup_rel n |> RelDecl.map_value (lift n) |> RelDecl.get_value |> Option.map EConstr.of_constr + env |> lookup_rel n |> RelDecl.get_value |> Option.map (EConstr.of_constr %> lift n) | _ -> None let special_red_case env sigma whfun (ci, p, c, lf) = - let open EConstr in let rec redrec s = let (constr, cargs) = whfun s in match match_eval_ref env sigma constr with @@ -614,7 +601,6 @@ let reduce_projection env sigma pb (recarg'hd,stack') stack = | _ -> NotReducible) let reduce_proj env sigma whfun whfun' c = - let open EConstr in let rec redrec s = match EConstr.kind sigma s with | Proj (proj, c) -> @@ -641,7 +627,7 @@ let whd_nothing_for_iota env sigma s = | Rel n -> let open Context.Rel.Declaration in (match lookup_rel n env with - | LocalDef (_,body,_) -> whrec (EConstr.of_constr (lift n body), stack) + | LocalDef (_,body,_) -> whrec (lift n (EConstr.of_constr body), stack) | _ -> s) | Var id -> let open Context.Named.Declaration in @@ -673,7 +659,6 @@ let whd_nothing_for_iota env sigma s = it fails if no redex is around *) let rec red_elim_const env sigma ref u largs = - let open EConstr in let nargs = List.length largs in let largs, unfold_anyway, unfold_nonelim, nocase = match recargs ref with @@ -747,7 +732,6 @@ and reduce_params env sigma stack l = a reducible iota/fix/cofix redex (the "simpl" tactic) *) and whd_simpl_stack env sigma = - let open EConstr in let rec redrec s = let (x, stack) = decompose_app_vect sigma s in let stack = Array.map_to_list EConstr.of_constr stack in @@ -818,7 +802,6 @@ and whd_simpl_stack env sigma = (* reduce until finding an applied constructor or fail *) and whd_construct_stack env sigma s = - let open EConstr in let (constr, cargs as s') = whd_simpl_stack env sigma s in if reducible_mind_case sigma constr then s' else match match_eval_ref env sigma constr with @@ -838,7 +821,6 @@ and whd_construct_stack env sigma s = let try_red_product env sigma c = let simpfun c = EConstr.of_constr (clos_norm_flags betaiotazeta env sigma c) in - let open EConstr in let rec redrec env x = let x = EConstr.of_constr (whd_betaiota sigma x) in match EConstr.kind sigma x with @@ -948,7 +930,6 @@ let whd_simpl_stack = immediately hides a non reducible fix or a cofix *) let whd_simpl_orelse_delta_but_fix env sigma c = - let open EConstr in let rec redrec s = let (constr, stack as s') = whd_simpl_stack env sigma s in match match_eval_ref_value env sigma constr with @@ -982,7 +963,6 @@ let simpl env sigma c = strong whd_simpl env sigma c (* Reduction at specific subterms *) let matches_head env sigma c t = - let open EConstr in match EConstr.kind sigma t with | App (f,_) -> Constr_matching.matches env sigma c f | Proj (p, _) -> Constr_matching.matches env sigma c (mkConstU (Projection.constant p, Univ.Instance.empty)) @@ -993,11 +973,9 @@ let matches_head env sigma c t = of the projection and its eta expanded form. *) let change_map_constr_with_binders_left_to_right g f (env, l as acc) sigma c = - let open EConstr in match EConstr.kind sigma c with | Proj (p, r) -> (* Treat specially for partial applications *) let t = Retyping.expand_projection env sigma p r [] in - let t = EConstr.of_constr t in let hdf, al = destApp sigma t in let a = al.(Array.length al - 1) in let app = (mkApp (hdf, Array.sub al 0 (Array.length al - 1))) in @@ -1011,7 +989,6 @@ let change_map_constr_with_binders_left_to_right g f (env, l as acc) sigma c = | _ -> map_constr_with_binders_left_to_right sigma g f acc c let e_contextually byhead (occs,c) f = { e_redfun = begin fun env sigma t -> - let open EConstr in let (nowhere_except_in,locs) = Locusops.convert_occs occs in let maxocc = List.fold_right max locs 0 in let pos = ref 1 in @@ -1138,7 +1115,6 @@ let unfoldn loccname env sigma c = (* Re-folding constants tactics: refold com in term c *) let fold_one_com com env sigma c = - let open EConstr in let rcom = try EConstr.of_constr (red_product env sigma com) with Redelimination -> error "Not reducible." in @@ -1176,7 +1152,6 @@ let compute = cbv_betadeltaiota * the specified occurrences. *) let abstract_scheme env sigma (locc,a) (c, sigma) = - let open EConstr in let ta = Retyping.get_type_of env sigma a in let ta = EConstr.of_constr ta in let na = named_hd env (EConstr.to_constr sigma ta) Anonymous in @@ -1189,7 +1164,6 @@ let abstract_scheme env sigma (locc,a) (c, sigma) = mkLambda (na,ta,c'), sigma' let pattern_occs loccs_trm = { e_redfun = begin fun env sigma c -> - let open EConstr in let sigma = Sigma.to_evar_map sigma in let abstr_trm, sigma = List.fold_right (abstract_scheme env sigma) loccs_trm (c,sigma) in try @@ -1218,7 +1192,6 @@ let check_not_primitive_record env ind = return name, B and t' *) let reduce_to_ind_gen allow_product env sigma t = - let open EConstr in let rec elimrec env t l = let t = hnf_constr env sigma t in let t = EConstr.of_constr t in @@ -1246,7 +1219,7 @@ let reduce_to_atomic_ind env sigma c = reduce_to_ind_gen false env sigma (EConst let find_hnf_rectype env sigma t = let ind,t = reduce_to_atomic_ind env sigma t in - ind, snd (decompose_app t) + ind, snd (Term.decompose_app t) (* Reduce the weak-head redex [beta,iota/fix/cofix[all],cast,zeta,simpl/delta] or raise [NotStepReducible] if not a weak-head redex *) @@ -1254,7 +1227,6 @@ let find_hnf_rectype env sigma t = exception NotStepReducible let one_step_reduce env sigma c = - let open EConstr in let rec redrec (x, stack) = match EConstr.kind sigma x with | Lambda (n,t,c) -> @@ -1302,7 +1274,6 @@ let reduce_to_ref_gen allow_product env sigma ref t = else (* lazily reduces to match the head of [t] with the expected [ref] *) let rec elimrec env t l = - let open EConstr in let c, _ = decompose_app_vect sigma t in let c = EConstr.of_constr c in match EConstr.kind sigma c with diff --git a/pretyping/typing.ml b/pretyping/typing.ml index 17adea5f2..cf58a0b66 100644 --- a/pretyping/typing.ml +++ b/pretyping/typing.ml @@ -65,7 +65,6 @@ let e_assumption_of_judgment env evdref j = error_assumption env !evdref j let e_judge_of_apply env evdref funj argjv = - let open EConstr in let rec apply_rec n typ = function | [] -> { uj_val = mkApp (j_val funj, Array.map j_val argjv); @@ -100,7 +99,6 @@ let max_sort l = if Sorts.List.mem InSet l then InSet else InProp let e_is_correct_arity env evdref c pj ind specif params = - let open EConstr in let arsign = make_arity_signature env true (make_ind_family (ind,params)) in let allowed_sorts = elim_sorts specif in let error () = Pretype_errors.error_elim_arity env !evdref ind allowed_sorts c pj None in @@ -124,7 +122,6 @@ let e_is_correct_arity env evdref c pj ind specif params = srec env pj.uj_type (List.rev arsign) let lambda_applist_assum sigma n c l = - let open EConstr in let rec app n subst t l = if Int.equal n 0 then if l == [] then substl subst t @@ -150,7 +147,6 @@ let e_type_case_branches env evdref (ind,largs) pj c = (lc, ty) let e_judge_of_case env evdref ci pj cj lfj = - let open EConstr in let indspec = try find_mrectype env !evdref cj.uj_type with Not_found -> error_case_not_inductive env !evdref cj in @@ -161,7 +157,6 @@ let e_judge_of_case env evdref ci pj cj lfj = uj_type = rslty } let check_type_fixpoint loc env evdref lna lar vdefj = - let open EConstr in let lt = Array.length vdefj in if Int.equal (Array.length lar) lt then for i = 0 to lt-1 do @@ -183,7 +178,6 @@ let check_allowed_sort env sigma ind c p = (Some(ksort,s,Type_errors.error_elim_explain ksort s)) let e_judge_of_cast env evdref cj k tj = - let open EConstr in let expected_type = tj.utj_val in if not (Evarconv.e_cumul env evdref cj.uj_type expected_type) then error_actual_type_core env !evdref cj expected_type; @@ -259,7 +253,6 @@ let judge_of_letin env name defj typj j = (* cstr must be in n.f. w.r.t. evars and execute returns a judgement where both the term and type are in n.f. *) let rec execute env evdref cstr = - let open EConstr in let cstr = EConstr.of_constr (whd_evar !evdref (EConstr.Unsafe.to_constr cstr)) in match EConstr.kind !evdref cstr with | Meta n -> diff --git a/pretyping/unification.ml b/pretyping/unification.ml index 1b209fa77..483fefaf2 100644 --- a/pretyping/unification.ml +++ b/pretyping/unification.ml @@ -85,7 +85,6 @@ let occur_meta_evd sigma mv c = gives [x1:A1]..[xn:An]c' such that c converts to ([x1:A1]..[xn:An]c' l) *) let abstract_scheme env evd c l lname_typ = - let open EConstr in let mkLambda_name env (n,a,b) = mkLambda (named_hd env (EConstr.Unsafe.to_constr a) n, a, b) in @@ -131,7 +130,6 @@ let set_occurrences_of_last_arg args = Some AllOccurrences :: List.tl (Array.map_to_list (fun _ -> None) args) let abstract_list_all_with_dependencies env evd typ c l = - let open EConstr in let evd = Sigma.Unsafe.of_evar_map evd in let Sigma (ev, evd, _) = new_evar env evd typ in let evd = Sigma.to_evar_map evd in @@ -195,8 +193,6 @@ let pose_all_metas_as_evars env evd t = (!evdref, c) let solve_pattern_eqn_array (env,nb) f l c (sigma,metasubst,evarsubst : subst0) = - let open EConstr in - let open Vars in match EConstr.kind sigma f with | Meta k -> (* We enforce that the Meta does not depend on the [nb] @@ -476,7 +472,6 @@ let use_evars_pattern_unification flags = && Flags.version_strictly_greater Flags.V8_2 let use_metas_pattern_unification sigma flags nb l = - let open EConstr in !global_pattern_unification_flag && flags.use_pattern_unification || (Flags.version_less_or_equal Flags.V8_3 || flags.use_meta_bound_pattern_unification) && @@ -636,7 +631,6 @@ let check_compatibility env pbty flags (sigma,metasubst,evarsubst : subst0) tyM let rec is_neutral env sigma ts t = - let open EConstr in let (f, l) = decompose_app_vect sigma t in match EConstr.kind sigma (EConstr.of_constr f) with | Const (c, u) -> @@ -666,7 +660,6 @@ let is_eta_constructor_app env sigma ts f l1 term = | _ -> false let eta_constructor_app env sigma f l1 term = - let open EConstr in match EConstr.kind sigma f with | Construct (((_, i as ind), j), u) -> let mib = lookup_mind (fst ind) env in @@ -684,8 +677,6 @@ let print_constr_env env c = print_constr_env env (EConstr.Unsafe.to_constr c) let rec unify_0_with_initial_metas (sigma,ms,es as subst : subst0) conv_at_top env cv_pb flags m n = - let open EConstr in - let open Vars in let rec unirec_rec (curenv,nb as curenvnb) pb opt ((sigma,metasubst,evarsubst) as substn : subst0) curm curn = let cM = Evarutil.whd_head_evar sigma curm and cN = Evarutil.whd_head_evar sigma curn in @@ -892,7 +883,7 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst : subst0) conv_at_top e let expand_proj c c' l = match EConstr.kind sigma c with | Proj (p, t) when not (Projection.unfolded p) && needs_expansion p c' -> - (try destApp sigma (EConstr.of_constr (Retyping.expand_projection curenv sigma p t (Array.to_list l))) + (try destApp sigma (Retyping.expand_projection curenv sigma p t (Array.to_list l)) with RetypeError _ -> (** Unification can be called on ill-typed terms, due to FO and eta in particular, fail gracefully in that case *) (c, l)) @@ -1128,8 +1119,6 @@ let right = false let rec unify_with_eta keptside flags env sigma c1 c2 = (* Question: try whd_all on ci if not two lambdas? *) - let open EConstr in - let open Vars in match EConstr.kind sigma c1, EConstr.kind sigma c2 with | (Lambda (na,t1,c1'), Lambda (_,t2,c2')) -> let env' = push_rel_assum (na,t1) env in @@ -1234,8 +1223,6 @@ let merge_instances env sigma flags st1 st2 c1 c2 = * since other metavars might also need to be resolved. *) let applyHead env (type r) (evd : r Sigma.t) n c = - let open EConstr in - let open Vars in let rec apprec : type s. _ -> _ -> _ -> (r, s) Sigma.le -> s Sigma.t -> (constr, r) Sigma.sigma = fun n c cty p evd -> if Int.equal n 0 then @@ -1307,7 +1294,6 @@ let order_metas metas = (* Solve an equation ?n[x1=u1..xn=un] = t where ?n is an evar *) let solve_simple_evar_eqn ts env evd ev rhs = - let open EConstr in match solve_simple_eqn (Evarconv.evar_conv_x ts) env evd (None,ev,rhs) with | UnifFailure (evd,reason) -> error_cannot_unify env evd ~reason (mkEvar ev,rhs); @@ -1319,8 +1305,6 @@ let solve_simple_evar_eqn ts env evd ev rhs = is true, unification of types of metas is required *) let w_merge env with_types flags (evd,metas,evars : subst0) = - let open EConstr in - let open Vars in let rec w_merge_rec evd metas evars eqns = (* Process evars *) @@ -1498,7 +1482,6 @@ let w_unify_core_0 env evd with_types cv_pb flags m n = let w_typed_unify env evd = w_unify_core_0 env evd true let w_typed_unify_array env evd flags f1 l1 f2 l2 = - let open EConstr in let f1 = EConstr.of_constr f1 in let f2 = EConstr.of_constr f2 in let l1 = Array.map EConstr.of_constr l1 in @@ -1529,7 +1512,6 @@ let iter_fail f a = contexts, with evars, and possibly with occurrences *) let indirectly_dependent sigma c d decls = - let open EConstr in not (isVar sigma c) && (* This test is not needed if the original term is a variable, but it is needed otherwise, as e.g. when abstracting over "2" in @@ -1590,7 +1572,6 @@ let default_matching_flags (sigma,_) = exception PatternNotFound let make_pattern_test from_prefix_of_ind is_correct_type env sigma (pending,c) = - let open EConstr in let flags = if from_prefix_of_ind then let flags = default_matching_flags pending in @@ -1600,7 +1581,6 @@ let make_pattern_test from_prefix_of_ind is_correct_type env sigma (pending,c) = else default_matching_flags pending in let n = Array.length (snd (decompose_app_vect sigma c)) in let matching_fun _ t = - let open EConstr in try let t',l2 = if from_prefix_of_ind then @@ -1754,8 +1734,6 @@ let keyed_unify env evd kop = Unifies [cl] to every subterm of [op] until it finds a match. Fails if no match is found *) let w_unify_to_subterm env evd ?(flags=default_unify_flags ()) (op,cl) = - let open EConstr in - let open Vars in let bestexn = ref None in let kop = Keys.constr_key (EConstr.to_constr evd op) in let rec matchrec cl = @@ -1831,8 +1809,6 @@ let w_unify_to_subterm env evd ?(flags=default_unify_flags ()) (op,cl) = Unifies [cl] to every subterm of [op] and return all the matches. Fails if no match is found *) let w_unify_to_subterm_all env evd ?(flags=default_unify_flags ()) (op,cl) = - let open EConstr in - let open Vars in let return a b = let (evd,c as a) = a () in if List.exists (fun (evd',c') -> EConstr.eq_constr evd' c c') b then b else a :: b @@ -1897,7 +1873,6 @@ let w_unify_to_subterm_all env evd ?(flags=default_unify_flags ()) (op,cl) = | _ -> res let w_unify_to_subterm_list env evd flags hdmeta oplist t = - let open EConstr in List.fold_right (fun op (evd,l) -> let op = whd_meta evd op in @@ -2008,7 +1983,6 @@ let w_unify2 env evd flags dep cv_pb ty1 ty2 = convertible and first-order otherwise. But if failed if e.g. the type of Meta(1) had meta-variables in it. *) let w_unify env evd cv_pb ?(flags=default_unify_flags ()) ty1 ty2 = - let open EConstr in let hd1,l1 = decompose_app_vect evd (EConstr.of_constr (whd_nored evd ty1)) in let hd2,l2 = decompose_app_vect evd (EConstr.of_constr (whd_nored evd ty2)) in let is_empty1 = Array.is_empty l1 in diff --git a/tactics/tactics.ml b/tactics/tactics.ml index 2cb9e0864..eebb2a038 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -3165,7 +3165,7 @@ let expand_projections env sigma c = let sigma = Sigma.to_evar_map sigma in let rec aux env c = match EConstr.kind sigma c with - | Proj (p, c) -> EConstr.of_constr (Retyping.expand_projection env sigma p (aux env c) []) + | Proj (p, c) -> Retyping.expand_projection env sigma p (aux env c) [] | _ -> map_constr_with_full_binders sigma push_rel aux env c in EConstr.Unsafe.to_constr (aux env (EConstr.of_constr c)) -- cgit v1.2.3 From 7267dfafe9215c35275a39814c8af451961e997c Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 11 Nov 2016 17:48:47 +0100 Subject: Goal API using EConstr. --- ide/ide_slave.ml | 2 +- plugins/decl_mode/decl_proof_instr.ml | 6 +++--- plugins/decl_mode/g_decl_mode.ml4 | 2 +- plugins/setoid_ring/newring.ml | 2 +- plugins/ssrmatching/ssrmatching.ml4 | 4 ++-- printing/printer.ml | 4 ++-- proofs/clenv.ml | 1 + proofs/goal.ml | 8 ++++++-- proofs/goal.mli | 10 +++++----- proofs/logic.ml | 16 ++++++++++------ proofs/tacmach.ml | 4 ++-- tactics/class_tactics.ml | 8 +++++--- tactics/hints.ml | 2 +- 13 files changed, 40 insertions(+), 29 deletions(-) diff --git a/ide/ide_slave.ml b/ide/ide_slave.ml index 48fd0a93e..c11c78587 100644 --- a/ide/ide_slave.ml +++ b/ide/ide_slave.ml @@ -191,7 +191,7 @@ let process_goal sigma g = let min_env = Environ.reset_context env in let id = Goal.uid g in let ccl = - let norm_constr = Reductionops.nf_evar sigma (Goal.V82.concl sigma g) in + let norm_constr = Reductionops.nf_evar sigma (EConstr.Unsafe.to_constr (Goal.V82.concl sigma g)) in Richpp.richpp_of_pp (pr_goal_concl_style_env env sigma norm_constr) in let process_hyp d (env,l) = diff --git a/plugins/decl_mode/decl_proof_instr.ml b/plugins/decl_mode/decl_proof_instr.ml index 44cd22c8b..3bb6f1b5d 100644 --- a/plugins/decl_mode/decl_proof_instr.ml +++ b/plugins/decl_mode/decl_proof_instr.ml @@ -41,7 +41,7 @@ let clear ids { it = goal; sigma } = let ids = List.fold_left (fun accu x -> Id.Set.add x accu) Id.Set.empty ids in let env = Goal.V82.env sigma goal in let sign = Goal.V82.hyps sigma goal in - let cl = Goal.V82.concl sigma goal in + let cl = EConstr.Unsafe.to_constr (Goal.V82.concl sigma goal) in let evdref = ref (Evd.clear_metas sigma) in let (hyps, concl) = try Evarutil.clear_hyps_in_evi env evdref sign cl ids @@ -49,7 +49,7 @@ let clear ids { it = goal; sigma } = user_err (str "Cannot clear " ++ pr_id id) in let sigma = !evdref in - let (gl,ev,sigma) = Goal.V82.mk_goal sigma hyps concl (Goal.V82.extra sigma goal) in + let (gl,ev,sigma) = Goal.V82.mk_goal sigma hyps (EConstr.of_constr concl) (Goal.V82.extra sigma goal) in let sigma = Goal.V82.partial_solution_to sigma goal gl ev in { it = [gl]; sigma } @@ -74,7 +74,7 @@ let tcl_change_info_gen info_gen = let concl = pf_concl gls in let hyps = Goal.V82.hyps (project gls) it in let extra = Goal.V82.extra (project gls) it in - let (gl,ev,sigma) = Goal.V82.mk_goal (project gls) hyps concl (info_gen extra) in + let (gl,ev,sigma) = Goal.V82.mk_goal (project gls) hyps (EConstr.of_constr concl) (info_gen extra) in let sigma = Goal.V82.partial_solution sigma it ev in { it = [gl] ; sigma= sigma; } ) diff --git a/plugins/decl_mode/g_decl_mode.ml4 b/plugins/decl_mode/g_decl_mode.ml4 index 18a35c6cf..9e2c9f597 100644 --- a/plugins/decl_mode/g_decl_mode.ml4 +++ b/plugins/decl_mode/g_decl_mode.ml4 @@ -25,7 +25,7 @@ open Ppdecl_proof let pr_goal gs = let (g,sigma) = Goal.V82.nf_evar (Tacmach.project gs) (Evd.sig_it gs) in let env = Goal.V82.env sigma g in - let concl = Goal.V82.concl sigma g in + let concl = EConstr.Unsafe.to_constr (Goal.V82.concl sigma g) in let goal = Printer.pr_context_of env sigma ++ cut () ++ str "============================" ++ cut () ++ diff --git a/plugins/setoid_ring/newring.ml b/plugins/setoid_ring/newring.ml index b0a3e839b..089e76d7a 100644 --- a/plugins/setoid_ring/newring.ml +++ b/plugins/setoid_ring/newring.ml @@ -172,7 +172,7 @@ let ltac_apply (f : Value.t) (args: Tacinterp.Value.t list) = let dummy_goal env sigma = let (gl,_,sigma) = - Goal.V82.mk_goal sigma (named_context_val env) mkProp Evd.Store.empty in + Goal.V82.mk_goal sigma (named_context_val env) EConstr.mkProp Evd.Store.empty in {Evd.it = gl; Evd.sigma = sigma} let constr_of v = match Value.to_constr v with diff --git a/plugins/ssrmatching/ssrmatching.ml4 b/plugins/ssrmatching/ssrmatching.ml4 index 308fb414e..7f628f165 100644 --- a/plugins/ssrmatching/ssrmatching.ml4 +++ b/plugins/ssrmatching/ssrmatching.ml4 @@ -1078,7 +1078,7 @@ END let thin id sigma goal = let ids = Id.Set.singleton id in let env = Goal.V82.env sigma goal in - let cl = Goal.V82.concl sigma goal in + let cl = EConstr.Unsafe.to_constr (Goal.V82.concl sigma goal) in let evdref = ref (Evd.clear_metas sigma) in let ans = try Some (Evarutil.clear_hyps_in_evi env evdref (Environ.named_context_val env) cl ids) @@ -1088,7 +1088,7 @@ let thin id sigma goal = | None -> sigma | Some (hyps, concl) -> let sigma = !evdref in - let (gl,ev,sigma) = Goal.V82.mk_goal sigma hyps concl (Goal.V82.extra sigma goal) in + let (gl,ev,sigma) = Goal.V82.mk_goal sigma hyps (EConstr.of_constr concl) (Goal.V82.extra sigma goal) in let sigma = Goal.V82.partial_solution_to sigma goal gl ev in sigma diff --git a/printing/printer.ml b/printing/printer.ml index b36df27ed..ed6ebdc9b 100644 --- a/printing/printer.ml +++ b/printing/printer.ml @@ -384,7 +384,7 @@ let pr_transparent_state (ids, csts) = let default_pr_goal gs = let (g,sigma) = Goal.V82.nf_evar (project gs) (sig_it gs) in let env = Goal.V82.env sigma g in - let concl = Goal.V82.concl sigma g in + let concl = EConstr.Unsafe.to_constr (Goal.V82.concl sigma g) in let goal = pr_context_of env sigma ++ cut () ++ str "============================" ++ cut () ++ @@ -407,7 +407,7 @@ let pr_goal_name sigma g = let pr_concl n sigma g = let (g,sigma) = Goal.V82.nf_evar sigma g in let env = Goal.V82.env sigma g in - let pc = pr_goal_concl_style_env env sigma (Goal.V82.concl sigma g) in + let pc = pr_goal_concl_style_env env sigma (EConstr.Unsafe.to_constr (Goal.V82.concl sigma g)) in str (emacs_str "") ++ str "subgoal " ++ int n ++ pr_goal_tag g ++ pr_goal_name sigma g ++ str " is:" ++ cut () ++ str" " ++ pc diff --git a/proofs/clenv.ml b/proofs/clenv.ml index 956be88c8..67ed5daaa 100644 --- a/proofs/clenv.ml +++ b/proofs/clenv.ml @@ -264,6 +264,7 @@ let clenv_unify_meta_types ?(flags=default_unify_flags ()) clenv = let clenv_unique_resolver ?(flags=default_unify_flags ()) clenv gl = let concl = Goal.V82.concl clenv.evd (sig_it gl) in + let concl = EConstr.Unsafe.to_constr concl in if isMeta (fst (decompose_appvect (whd_nored clenv.evd (EConstr.of_constr clenv.templtyp.rebus)))) then clenv_unify CUMUL ~flags (clenv_type clenv) concl (clenv_unify_meta_types ~flags clenv) diff --git a/proofs/goal.ml b/proofs/goal.ml index bcb14e2d6..4c598ca29 100644 --- a/proofs/goal.ml +++ b/proofs/goal.ml @@ -62,6 +62,7 @@ module V82 = struct be shelved. It must not appear as a future_goal, so the future goals are restored to their initial value after the evar is created. *) + let concl = EConstr.Unsafe.to_constr concl in let prev_future_goals = Evd.future_goals evars in let prev_principal_goal = Evd.principal_future_goal evars in let evi = { Evd.evar_hyps = hyps; @@ -78,13 +79,14 @@ module V82 = struct let evars = Sigma.to_evar_map evars in let evars = Evd.restore_future_goals evars prev_future_goals prev_principal_goal in let ctxt = Environ.named_context_of_val hyps in - let inst = Array.map_of_list (NamedDecl.get_id %> mkVar) ctxt in - let ev = Term.mkEvar (evk,inst) in + let inst = Array.map_of_list (NamedDecl.get_id %> EConstr.mkVar) ctxt in + let ev = EConstr.mkEvar (evk,inst) in (evk, ev, evars) (* Instantiates a goal with an open term *) let partial_solution sigma evk c = (* Check that the goal itself does not appear in the refined term *) + let c = EConstr.Unsafe.to_constr c in let _ = if not (Evarutil.occur_evar_upto sigma evk c) then () else Pretype_errors.error_occur_check Environ.empty_env sigma evk (EConstr.of_constr c) @@ -158,4 +160,6 @@ module V82 = struct t ) ~init:(concl sigma gl) env + let concl sigma gl = EConstr.of_constr (concl sigma gl) + end diff --git a/proofs/goal.mli b/proofs/goal.mli index 6a79c1f45..ca6584e76 100644 --- a/proofs/goal.mli +++ b/proofs/goal.mli @@ -38,7 +38,7 @@ module V82 : sig val nf_hyps : Evd.evar_map -> goal -> Environ.named_context_val (* Access to ".evar_concl" *) - val concl : Evd.evar_map -> goal -> Term.constr + val concl : Evd.evar_map -> goal -> EConstr.constr (* Access to ".evar_extra" *) val extra : Evd.evar_map -> goal -> Evd.Store.t @@ -48,16 +48,16 @@ module V82 : sig the evar corresponding to the goal, and an updated evar_map. *) val mk_goal : Evd.evar_map -> Environ.named_context_val -> - Term.constr -> + EConstr.constr -> Evd.Store.t -> - goal * Term.constr * Evd.evar_map + goal * EConstr.constr * Evd.evar_map (* Instantiates a goal with an open term *) - val partial_solution : Evd.evar_map -> goal -> Term.constr -> Evd.evar_map + val partial_solution : Evd.evar_map -> goal -> EConstr.constr -> Evd.evar_map (* Instantiates a goal with an open term, reusing name of goal for second goal *) - val partial_solution_to : Evd.evar_map -> goal -> goal -> Term.constr -> Evd.evar_map + val partial_solution_to : Evd.evar_map -> goal -> goal -> EConstr.constr -> Evd.evar_map (* Principal part of the weak-progress tactical *) val weak_progress : goal list Evd.sigma -> goal Evd.sigma -> bool diff --git a/proofs/logic.ml b/proofs/logic.ml index 093d949d5..c5d6e3e08 100644 --- a/proofs/logic.ml +++ b/proofs/logic.ml @@ -349,7 +349,8 @@ let rec mk_refgoals sigma goal goalacc conclty trm = let conclty = nf_betaiota sigma (EConstr.of_constr conclty) in if !check && occur_meta sigma (EConstr.of_constr conclty) then raise (RefinerError (MetaInType conclty)); - let (gl,ev,sigma) = mk_goal hyps conclty in + let (gl,ev,sigma) = mk_goal hyps (EConstr.of_constr conclty) in + let ev = EConstr.Unsafe.to_constr ev in gl::goalacc, conclty, sigma, ev | Cast (t,k, ty) -> @@ -424,7 +425,8 @@ and mk_hdgoals sigma goal goalacc trm = match kind_of_term trm with | Cast (c,_, ty) when isMeta c -> check_typability env sigma ty; - let (gl,ev,sigma) = mk_goal hyps (nf_betaiota sigma (EConstr.of_constr ty)) in + let (gl,ev,sigma) = mk_goal hyps (EConstr.of_constr (nf_betaiota sigma (EConstr.of_constr ty))) in + let ev = EConstr.Unsafe.to_constr ev in gl::goalacc,ty,sigma,ev | Cast (t,_, ty) -> @@ -526,6 +528,7 @@ let prim_refiner r sigma goal = let env = Goal.V82.env sigma goal in let sign = Goal.V82.hyps sigma goal in let cl = Goal.V82.concl sigma goal in + let cl = EConstr.Unsafe.to_constr cl in let mk_goal hyps concl = Goal.V82.mk_goal sigma hyps concl (Goal.V82.extra sigma goal) in @@ -533,7 +536,7 @@ let prim_refiner r sigma goal = (* Logical rules *) | Cut (b,replace,id,t) -> (* if !check && not (Retyping.get_sort_of env sigma t) then*) - let (sg1,ev1,sigma) = mk_goal sign (nf_betaiota sigma (EConstr.of_constr t)) in + let (sg1,ev1,sigma) = mk_goal sign (EConstr.of_constr (nf_betaiota sigma (EConstr.of_constr t))) in let sign,t,cl,sigma = if replace then let nexthyp = get_hyp_after id (named_context_of_val sign) in @@ -546,9 +549,10 @@ let prim_refiner r sigma goal = user_err ~hdr:"Logic.prim_refiner" (str "Variable " ++ pr_id id ++ str " is already declared."); push_named_context_val (LocalAssum (id,t)) sign,t,cl,sigma) in + let t = EConstr.of_constr t in let (sg2,ev2,sigma) = - Goal.V82.mk_goal sigma sign cl (Goal.V82.extra sigma goal) in - let oterm = Term.mkNamedLetIn id ev1 t ev2 in + Goal.V82.mk_goal sigma sign (EConstr.of_constr cl) (Goal.V82.extra sigma goal) in + let oterm = EConstr.mkLetIn (Name id, ev1, t, EConstr.Vars.subst_var id ev2) in let sigma = Goal.V82.partial_solution_to sigma goal sg2 oterm in if b then ([sg1;sg2],sigma) else ([sg2;sg1],sigma) @@ -556,5 +560,5 @@ let prim_refiner r sigma goal = check_meta_variables c; let (sgl,cl',sigma,oterm) = mk_refgoals sigma goal [] cl c in let sgl = List.rev sgl in - let sigma = Goal.V82.partial_solution sigma goal oterm in + let sigma = Goal.V82.partial_solution sigma goal (EConstr.of_constr oterm) in (sgl, sigma) diff --git a/proofs/tacmach.ml b/proofs/tacmach.ml index 5e9c0ba8e..c2ebb76d7 100644 --- a/proofs/tacmach.ml +++ b/proofs/tacmach.ml @@ -41,7 +41,7 @@ let project = Refiner.project let pf_env = Refiner.pf_env let pf_hyps = Refiner.pf_hyps -let pf_concl gls = Goal.V82.concl (project gls) (sig_it gls) +let pf_concl gls = EConstr.Unsafe.to_constr (Goal.V82.concl (project gls) (sig_it gls)) let pf_hyps_types gls = let sign = Environ.named_context (pf_env gls) in List.map (function LocalAssum (id,x) @@ -137,7 +137,7 @@ open Pp let db_pr_goal sigma g = let env = Goal.V82.env sigma g in let penv = print_named_context env in - let pc = print_constr_env env (Goal.V82.concl sigma g) in + let pc = print_constr_env env (EConstr.Unsafe.to_constr (Goal.V82.concl sigma g)) in str" " ++ hv 0 (penv ++ fnl () ++ str "============================" ++ fnl () ++ str" " ++ pc) ++ fnl () diff --git a/tactics/class_tactics.ml b/tactics/class_tactics.ml index bc1d0ed6b..be8d7eaa5 100644 --- a/tactics/class_tactics.ml +++ b/tactics/class_tactics.ml @@ -185,7 +185,7 @@ let set_typeclasses_depth = let pr_ev evs ev = Printer.pr_constr_env (Goal.V82.env evs ev) evs - (Evarutil.nf_evar evs (Goal.V82.concl evs ev)) + (Evarutil.nf_evar evs (EConstr.Unsafe.to_constr (Goal.V82.concl evs ev))) (** Typeclasses instance search tactic / eauto *) @@ -672,6 +672,7 @@ module V85 = struct let hints_tac hints sk fk {it = gl,info; sigma = s} = let env = Goal.V82.env s gl in let concl = Goal.V82.concl s gl in + let concl = EConstr.Unsafe.to_constr concl in let tacgl = {it = gl; sigma = s;} in let secvars = secvars_of_hyps (Environ.named_context_of_val (Goal.V82.hyps s gl)) in let poss = e_possible_resolve hints info.hints secvars info.only_classes s concl in @@ -784,7 +785,7 @@ module V85 = struct let fk'' = if not info.unique && List.is_empty gls' && not (needs_backtrack (Goal.V82.env s gl) s - info.is_evar (Goal.V82.concl s gl)) + info.is_evar (EConstr.Unsafe.to_constr (Goal.V82.concl s gl))) then fk else fk' in @@ -1458,7 +1459,7 @@ let _ = let resolve_one_typeclass env ?(sigma=Evd.empty) gl unique = let nc, gl, subst, _, _ = Evarutil.push_rel_context_to_named_context env gl in let (gl,t,sigma) = - Goal.V82.mk_goal sigma nc (EConstr.Unsafe.to_constr gl) Store.empty in + Goal.V82.mk_goal sigma nc gl Store.empty in let gls = { it = gl ; sigma = sigma; } in let hints = searchtable_map typeclasses_db in let st = Hint_db.transparent_state hints in @@ -1473,6 +1474,7 @@ let resolve_one_typeclass env ?(sigma=Evd.empty) gl unique = with Refiner.FailError _ -> raise Not_found in let evd = sig_sig gls' in + let t = EConstr.Unsafe.to_constr t in let t' = let (ev, inst) = destEvar t in mkEvar (ev, Array.of_list subst) in diff --git a/tactics/hints.ml b/tactics/hints.ml index e8225df2d..57358bb76 100644 --- a/tactics/hints.ml +++ b/tactics/hints.ml @@ -1423,7 +1423,7 @@ let pr_applicable_hint () = match glss.Evd.it with | [] -> CErrors.error "No focused goal." | g::_ -> - pr_hint_term glss.Evd.sigma (Goal.V82.concl glss.Evd.sigma g) + pr_hint_term glss.Evd.sigma (EConstr.Unsafe.to_constr (Goal.V82.concl glss.Evd.sigma g)) let pp_hint_mode = function | ModeInput -> str"+" -- cgit v1.2.3 From 53fe23265daafd47e759e73e8f97361c7fdd331b Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 11 Nov 2016 18:20:29 +0100 Subject: Refine API using EConstr. --- engine/evarutil.ml | 1 + engine/evarutil.mli | 2 +- ltac/extratactics.ml4 | 5 ++++- ltac/rewrite.ml | 6 +++--- proofs/goal.ml | 5 ++--- proofs/refine.ml | 11 ++++++----- proofs/refine.mli | 8 ++++---- tactics/class_tactics.ml | 2 +- tactics/inv.ml | 1 + tactics/tactics.ml | 42 +++++++++++++++++++++++++++++------------- toplevel/classes.ml | 2 +- 11 files changed, 53 insertions(+), 32 deletions(-) diff --git a/engine/evarutil.ml b/engine/evarutil.ml index bc55ac458..7ccf9d810 100644 --- a/engine/evarutil.ml +++ b/engine/evarutil.ml @@ -719,6 +719,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) diff --git a/engine/evarutil.mli b/engine/evarutil.mli index 8f3c3c352..431d98083 100644 --- a/engine/evarutil.mli +++ b/engine/evarutil.mli @@ -128,7 +128,7 @@ 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 -> EConstr.t -> bool (** {6 Value/Type constraints} *) diff --git a/ltac/extratactics.ml4 b/ltac/extratactics.ml4 index 981ff549d..c3ca8f906 100644 --- a/ltac/extratactics.ml4 +++ b/ltac/extratactics.ml4 @@ -353,7 +353,10 @@ let refine_tac ist simple c = let flags = constr_flags in let expected_type = Pretyping.OfType (EConstr.of_constr concl) in let c = Pretyping.type_uconstr ~flags ~expected_type ist c in - let update = { run = fun sigma -> c.delayed env sigma } in + let update = { run = fun sigma -> + let Sigma (c, sigma, p) = c.delayed env sigma in + Sigma (EConstr.of_constr c, sigma, p) + } in let refine = Refine.refine ~unsafe:true update in if simple then refine else refine <*> diff --git a/ltac/rewrite.ml b/ltac/rewrite.ml index 4d65b4c69..dfcfbfbd3 100644 --- a/ltac/rewrite.ml +++ b/ltac/rewrite.ml @@ -1556,7 +1556,7 @@ let assert_replacing id newt tac = if Id.equal n id then ev' else mkVar n in let (e, _) = destEvar ev in - Sigma (mkEvar (e, Array.map_of_list map nc), sigma, p +> q) + Sigma (EConstr.of_constr (mkEvar (e, Array.map_of_list map nc)), sigma, p +> q) end } end } in Proofview.tclTHEN prf (Proofview.tclFOCUS 2 2 tac) @@ -1582,7 +1582,7 @@ let cl_rewrite_clause_newtac ?abs ?origsigma ~progress strat clause = match clause, prf with | Some id, Some p -> let tac = tclTHENLIST [ - Refine.refine ~unsafe:false { run = fun h -> Sigma.here p h }; + Refine.refine ~unsafe:false { run = fun h -> Sigma.here (EConstr.of_constr p) h }; Proofview.Unsafe.tclNEWGOALS gls; ] in Proofview.Unsafe.tclEVARS undef <*> @@ -1597,7 +1597,7 @@ let cl_rewrite_clause_newtac ?abs ?origsigma ~progress strat clause = let env = Proofview.Goal.env gl in let make = { run = begin fun sigma -> let Sigma (ev, sigma, q) = Evarutil.new_evar env sigma (EConstr.of_constr newt) in - Sigma (mkApp (p, [| ev |]), sigma, q) + Sigma (EConstr.of_constr (mkApp (p, [| ev |])), sigma, q) end } in Refine.refine ~unsafe:false make <*> Proofview.Unsafe.tclNEWGOALS gls end } diff --git a/proofs/goal.ml b/proofs/goal.ml index 4c598ca29..7499bc251 100644 --- a/proofs/goal.ml +++ b/proofs/goal.ml @@ -86,12 +86,11 @@ module V82 = struct (* Instantiates a goal with an open term *) let partial_solution sigma evk c = (* Check that the goal itself does not appear in the refined term *) - let c = EConstr.Unsafe.to_constr c in let _ = if not (Evarutil.occur_evar_upto sigma evk c) then () - else Pretype_errors.error_occur_check Environ.empty_env sigma evk (EConstr.of_constr c) + else Pretype_errors.error_occur_check Environ.empty_env sigma evk c in - Evd.define evk c sigma + Evd.define evk (EConstr.Unsafe.to_constr c) sigma (* Instantiates a goal with an open term, using name of goal for evk' *) let partial_solution_to sigma evk evk' c = diff --git a/proofs/refine.ml b/proofs/refine.ml index 067764c00..11eabe0a9 100644 --- a/proofs/refine.ml +++ b/proofs/refine.ml @@ -47,7 +47,7 @@ let typecheck_evar ev env sigma = let typecheck_proof c concl env sigma = let evdref = ref sigma in - let () = Typing.e_check env evdref (EConstr.of_constr c) (EConstr.of_constr concl) in + let () = Typing.e_check env evdref c concl in !evdref let (pr_constrv,pr_constr) = @@ -77,6 +77,7 @@ let make_refine_enter ?(unsafe = true) f = let sigma = Sigma.to_evar_map sigma in let env = Proofview.Goal.env gl in let concl = Proofview.Goal.concl gl in + let concl = EConstr.of_constr concl in (** Save the [future_goals] state to restore them after the refinement. *) let prev_future_goals = Evd.future_goals sigma in @@ -98,9 +99,10 @@ let make_refine_enter ?(unsafe = true) f = let self = Proofview.Goal.goal gl in let _ = if not (Evarutil.occur_evar_upto sigma self c) then () - else Pretype_errors.error_occur_check env sigma self (EConstr.of_constr c) + else Pretype_errors.error_occur_check env sigma self c in (** Proceed to the refinement *) + let c = EConstr.Unsafe.to_constr c in let sigma = match evkmain with | None -> Evd.define self c sigma | Some evk -> @@ -133,23 +135,22 @@ let refine ?(unsafe = true) f = (** Useful definitions *) let with_type env evd c t = - let c = EConstr.of_constr c in let my_type = Retyping.get_type_of env evd c in let my_type = EConstr.of_constr my_type in let j = Environ.make_judge c my_type in let (evd,j') = - Coercion.inh_conv_coerce_to true (Loc.ghost) env evd j (EConstr.of_constr t) + Coercion.inh_conv_coerce_to true (Loc.ghost) env evd j t in evd , j'.Environ.uj_val let refine_casted ?unsafe f = Proofview.Goal.enter { enter = begin fun gl -> let gl = Proofview.Goal.assume gl in let concl = Proofview.Goal.concl gl in + let concl = EConstr.of_constr concl in let env = Proofview.Goal.env gl in let f = { run = fun h -> let Sigma (c, h, p) = f.run h in let sigma, c = with_type env (Sigma.to_evar_map h) c concl in - let c = EConstr.Unsafe.to_constr c in Sigma (c, Sigma.Unsafe.of_evar_map sigma, p) } in refine ?unsafe f diff --git a/proofs/refine.mli b/proofs/refine.mli index 4158e446c..205b97974 100644 --- a/proofs/refine.mli +++ b/proofs/refine.mli @@ -21,7 +21,7 @@ val pr_constr : (** {7 Refinement primitives} *) -val refine : ?unsafe:bool -> Constr.t Sigma.run -> unit tactic +val refine : ?unsafe:bool -> EConstr.t Sigma.run -> unit tactic (** In [refine ?unsafe t], [t] is a term with holes under some [evar_map] context. The term [t] is used as a partial solution for the current goal (refine is a goal-dependent tactic), the @@ -30,16 +30,16 @@ val refine : ?unsafe:bool -> Constr.t Sigma.run -> unit tactic tactic failures. If [unsafe] is [false] (default is [true]) [t] is type-checked beforehand. *) -val refine_one : ?unsafe:bool -> ('a * Constr.t) Sigma.run -> 'a tactic +val refine_one : ?unsafe:bool -> ('a * EConstr.t) Sigma.run -> 'a tactic (** A generalization of [refine] which assumes exactly one goal under focus *) (** {7 Helper functions} *) val with_type : Environ.env -> Evd.evar_map -> - Term.constr -> Term.types -> Evd.evar_map * EConstr.constr + EConstr.constr -> EConstr.types -> Evd.evar_map * EConstr.constr (** [with_type env sigma c t] ensures that [c] is of type [t] inserting a coercion if needed. *) -val refine_casted : ?unsafe:bool -> Constr.t Sigma.run -> unit tactic +val refine_casted : ?unsafe:bool -> EConstr.t Sigma.run -> unit tactic (** Like {!refine} except the refined term is coerced to the conclusion of the current goal. *) diff --git a/tactics/class_tactics.ml b/tactics/class_tactics.ml index be8d7eaa5..b0f355170 100644 --- a/tactics/class_tactics.ml +++ b/tactics/class_tactics.ml @@ -269,7 +269,7 @@ let unify_resolve_refine poly flags = {Environ.uj_val = term; Environ.uj_type = cl.cl_concl} concl; !evdref - in Sigma.here term (Sigma.Unsafe.of_evar_map sigma') } + in Sigma.here (EConstr.of_constr term) (Sigma.Unsafe.of_evar_map sigma') } end } (** Dealing with goals of the form A -> B and hints of the form diff --git a/tactics/inv.ml b/tactics/inv.ml index 9324d8e37..eebc67222 100644 --- a/tactics/inv.ml +++ b/tactics/inv.ml @@ -458,6 +458,7 @@ let raw_inversion inv_kind id status names = in let refined id = let prf = mkApp (mkVar id, args) in + let prf = EConstr.of_constr prf in Refine.refine { run = fun h -> Sigma (prf, h, Sigma.refl) } in let neqns = List.length realargs in diff --git a/tactics/tactics.ml b/tactics/tactics.ml index eebb2a038..639a12b34 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -177,7 +177,7 @@ let unsafe_intro env store decl b = let ninst = mkRel 1 :: inst in let nb = subst1 (mkVar (NamedDecl.get_id decl)) b in let Sigma (ev, sigma, p) = new_evar_instance nctx sigma (EConstr.of_constr nb) ~principal:true ~store ninst in - Sigma (mkNamedLambda_or_LetIn decl ev, sigma, p) + Sigma (EConstr.of_constr (mkNamedLambda_or_LetIn decl ev), sigma, p) end } let introduction ?(check=true) id = @@ -218,7 +218,7 @@ let convert_concl ?(check=true) ty k = end else Sigma.here () sigma in let Sigma (x, sigma, q) = Evarutil.new_evar env sigma ~principal:true ~store ty in let ans = if k == DEFAULTcast then x else mkCast(x,k,conclty) in - Sigma (ans, sigma, p +> q) + Sigma (EConstr.of_constr ans, sigma, p +> q) end } end } @@ -231,7 +231,8 @@ let convert_hyp ?(check=true) d = let sign = convert_hyp check (named_context_val env) sigma d in let env = reset_with_named_context sign env in Refine.refine ~unsafe:true { run = begin fun sigma -> - Evarutil.new_evar env sigma ~principal:true ~store (EConstr.of_constr ty) + let Sigma (c, sigma, p) = Evarutil.new_evar env sigma ~principal:true ~store (EConstr.of_constr ty) in + Sigma (EConstr.of_constr c, sigma, p) end } end } @@ -301,7 +302,8 @@ let clear_gen fail = function in let env = reset_with_named_context hyps env in let tac = Refine.refine ~unsafe:true { run = fun sigma -> - Evarutil.new_evar env sigma ~principal:true (EConstr.of_constr concl) + let Sigma (c, sigma, p) = Evarutil.new_evar env sigma ~principal:true (EConstr.of_constr concl) in + Sigma (EConstr.of_constr c, sigma, p) } in Sigma.Unsafe.of_pair (tac, !evdref) end } @@ -331,7 +333,8 @@ let move_hyp id dest = let sign' = move_hyp_in_named_context sigma id dest sign in let env = reset_with_named_context sign' env in Refine.refine ~unsafe:true { run = begin fun sigma -> - Evarutil.new_evar env sigma ~principal:true ~store (EConstr.of_constr ty) + let Sigma (c, sigma, p) = Evarutil.new_evar env sigma ~principal:true ~store (EConstr.of_constr ty) in + Sigma (EConstr.of_constr c, sigma, p) end } end } @@ -385,7 +388,8 @@ let rename_hyp repl = let nctx = Environ.val_of_named_context nhyps in let instance = List.map (NamedDecl.get_id %> mkVar) hyps in Refine.refine ~unsafe:true { run = begin fun sigma -> - Evarutil.new_evar_instance nctx sigma (EConstr.of_constr nconcl) ~principal:true ~store instance + let Sigma (c, sigma, p) = Evarutil.new_evar_instance nctx sigma (EConstr.of_constr nconcl) ~principal:true ~store instance in + Sigma (EConstr.of_constr c, sigma, p) end } end } @@ -541,6 +545,7 @@ let mutual_fix f n rest j = Proofview.Goal.nf_enter { enter = begin fun gl -> let typarray = Array.of_list (List.map pi3 all) in let bodies = Array.of_list evs in let oterm = Term.mkFix ((indxs,0),(funnames,typarray,bodies)) in + let oterm = EConstr.of_constr oterm in Sigma (oterm, sigma, p) end } end } @@ -592,6 +597,7 @@ let mutual_cofix f others j = Proofview.Goal.nf_enter { enter = begin fun gl -> let typarray = Array.of_list types in let bodies = Array.of_list evs in let oterm = Term.mkCoFix (0, (funnames, typarray, bodies)) in + let oterm = EConstr.of_constr oterm in Sigma (oterm, sigma, p) end } end } @@ -1248,6 +1254,7 @@ let cut c = let Sigma (f, h, p) = Evarutil.new_evar ~principal:true env h (EConstr.of_constr (mkArrow c (Vars.lift 1 concl))) in let Sigma (x, h, q) = Evarutil.new_evar env h (EConstr.of_constr c) in let f = mkLetIn (Name id, x, c, mkApp (Vars.lift 1 f, [|mkRel 1|])) in + let f = EConstr.of_constr f in Sigma (f, h, p +> q) end } else @@ -1680,6 +1687,7 @@ let solve_remaining_apply_goals = let concl = EConstr.of_constr concl in if Typeclasses.is_class_type evd concl then let evd', c' = Typeclasses.resolve_one_typeclass env evd concl in + let c' = EConstr.of_constr c' in let tac = Refine.refine ~unsafe:true { run = fun h -> Sigma.here c' h } in Sigma.Unsafe.of_pair (tac, evd') else Sigma.here (Proofview.tclUNIT ()) sigma @@ -1921,6 +1929,7 @@ let cut_and_apply c = let Sigma (f, sigma, p) = Evarutil.new_evar env sigma (EConstr.of_constr typ) in let Sigma (x, sigma, q) = Evarutil.new_evar env sigma (EConstr.of_constr c1) in let ans = mkApp (f, [|mkApp (c, [|x|])|]) in + let ans = EConstr.of_constr ans in Sigma (ans, sigma, p +> q) end } | _ -> error "lapply needs a non-dependent product." @@ -1937,6 +1946,7 @@ let cut_and_apply c = (* let refine_no_check = Profile.profile2 refine_no_checkkey refine_no_check *) let exact_no_check c = + let c = EConstr.of_constr c in Refine.refine ~unsafe:true { run = fun h -> Sigma.here c h } let exact_check c = @@ -1968,6 +1978,7 @@ let exact_proof c = Refine.refine { run = begin fun sigma -> let sigma = Sigma.to_evar_map sigma in let (c, ctx) = Constrintern.interp_casted_constr (pf_env gl) sigma c (pf_concl gl) in + let c = EConstr.of_constr c in let sigma = Evd.merge_universe_context sigma ctx in Sigma.Unsafe.of_pair (c, sigma) end } @@ -2083,7 +2094,8 @@ let clear_body ids = in check <*> Refine.refine ~unsafe:true { run = begin fun sigma -> - Evarutil.new_evar env sigma ~principal:true (EConstr.of_constr concl) + let Sigma (c, sigma, p) = Evarutil.new_evar env sigma ~principal:true (EConstr.of_constr concl) in + Sigma (EConstr.of_constr c, sigma, p) end } end } @@ -2137,7 +2149,7 @@ let apply_type newcl args = let newcl = nf_betaiota (Sigma.to_evar_map sigma) (EConstr.of_constr newcl) (* As in former Logic.refine *) in let Sigma (ev, sigma, p) = Evarutil.new_evar env sigma ~principal:true ~store (EConstr.of_constr newcl) in - Sigma (applist (ev, args), sigma, p) + Sigma (EConstr.of_constr (applist (ev, args)), sigma, p) end } end } @@ -2157,7 +2169,7 @@ let bring_hyps hyps = Refine.refine { run = begin fun sigma -> let Sigma (ev, sigma, p) = Evarutil.new_evar env sigma ~principal:true ~store (EConstr.of_constr newcl) in - Sigma (mkApp (ev, args), sigma, p) + Sigma (EConstr.of_constr (mkApp (ev, args)), sigma, p) end } end } @@ -2683,11 +2695,11 @@ let mkletin_goal env sigma store with_eq dep (id,lastlhyp,ccl,c) ty = let refl = applist (refl, [t;mkVar id]) in let newenv = insert_before [LocalAssum (heq,eq); decl] lastlhyp env in let Sigma (x, sigma, r) = new_evar newenv sigma ~principal:true ~store (EConstr.of_constr ccl) in - Sigma (mkNamedLetIn id c t (mkNamedLetIn heq refl eq x), sigma, p +> q +> r) + Sigma (EConstr.of_constr (mkNamedLetIn id c t (mkNamedLetIn heq refl eq x)), sigma, p +> q +> r) | None -> let newenv = insert_before [decl] lastlhyp env in let Sigma (x, sigma, p) = new_evar newenv sigma ~principal:true ~store (EConstr.of_constr ccl) in - Sigma (mkNamedLetIn id c t x, sigma, p) + Sigma (EConstr.of_constr (mkNamedLetIn id c t x), sigma, p) let letin_tac with_eq id c ty occs = Proofview.Goal.nf_s_enter { s_enter = begin fun gl -> @@ -2875,7 +2887,7 @@ let new_generalize_gen_let lconstr = let tac = Refine.refine { run = begin fun sigma -> let Sigma (ev, sigma, p) = Evarutil.new_evar env sigma ~principal:true (EConstr.of_constr newcl) in - Sigma ((applist (ev, args)), sigma, p) + Sigma (EConstr.of_constr (applist (ev, args)), sigma, p) end } in Sigma.Unsafe.of_pair (tac, sigma) @@ -3569,7 +3581,7 @@ let make_abstract_generalize env id typ concl dep ctx body c eqs args refls = (* Apply the reflexivity proofs on the indices. *) let appeqs = mkApp (instc, Array.of_list refls) in (* Finally, apply the reflexivity proof for the original hyp, to get a term of type gl again. *) - Sigma (mkApp (appeqs, abshypt), sigma, p) + Sigma (EConstr.of_constr (mkApp (appeqs, abshypt)), sigma, p) end } let hyps_of_vars env sigma sign nogen hyps = @@ -5005,6 +5017,10 @@ module New = struct {onhyps=None; concl_occs=AllOccurrences } let refine ?unsafe c = + let c = { run = begin fun sigma -> + let Sigma (c, sigma, p) = c.run sigma in + Sigma (EConstr.of_constr c, sigma, p) + end } in Refine.refine ?unsafe c <*> reduce_after_refine end diff --git a/toplevel/classes.ml b/toplevel/classes.ml index 1055d28b6..7c4ad6225 100644 --- a/toplevel/classes.ml +++ b/toplevel/classes.ml @@ -332,7 +332,7 @@ let new_instance ?(abstract=false) ?(global=false) ?(refine= !refine_instance) p if not (Option.is_empty term) then let init_refine = Tacticals.New.tclTHENLIST [ - Refine.refine { run = fun evm -> Sigma (Option.get term, evm, Sigma.refl) }; + Refine.refine { run = fun evm -> Sigma (EConstr.of_constr (Option.get term), evm, Sigma.refl) }; Proofview.Unsafe.tclNEWGOALS gls; Tactics.New.reduce_after_refine; ] -- cgit v1.2.3 From cbea91d815f134d63d02d8fb1bd78ed97db28cd1 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 11 Nov 2016 19:52:48 +0100 Subject: Tacmach API using EConstr. --- ltac/extratactics.ml4 | 4 +-- ltac/rewrite.ml | 2 +- plugins/btauto/refl_btauto.ml | 1 + plugins/cc/ccalgo.ml | 4 +-- plugins/cc/cctac.ml | 11 +++--- plugins/decl_mode/decl_proof_instr.ml | 14 ++++---- plugins/firstorder/instances.ml | 2 +- plugins/firstorder/sequent.ml | 4 +-- plugins/funind/functional_principles_proofs.ml | 22 +++++++----- plugins/funind/indfun.ml | 4 +-- plugins/funind/invfun.ml | 12 +++---- plugins/funind/recdef.ml | 12 +++---- plugins/omega/coq_omega.ml | 18 +++++----- plugins/romega/const_omega.ml | 2 +- plugins/ssrmatching/ssrmatching.ml4 | 2 +- pretyping/tacred.ml | 4 +-- pretyping/tacred.mli | 6 ++-- proofs/clenvtac.ml | 2 +- proofs/logic.ml | 1 + proofs/tacmach.ml | 45 ++++++++++++----------- proofs/tacmach.mli | 50 +++++++++++++------------- tactics/class_tactics.ml | 4 +-- tactics/contradiction.ml | 2 +- tactics/eauto.ml | 2 +- tactics/elim.ml | 4 +-- tactics/eqdecide.ml | 4 +-- tactics/equality.ml | 15 ++++---- tactics/hipattern.ml | 3 +- tactics/inv.ml | 2 +- tactics/tacticals.ml | 4 +-- tactics/tactics.ml | 44 +++++++++++++++-------- toplevel/auto_ind_decl.ml | 4 +-- toplevel/command.ml | 4 +-- 33 files changed, 170 insertions(+), 144 deletions(-) diff --git a/ltac/extratactics.ml4 b/ltac/extratactics.ml4 index c3ca8f906..3e7cf5d13 100644 --- a/ltac/extratactics.ml4 +++ b/ltac/extratactics.ml4 @@ -734,7 +734,7 @@ let refl_equal = call it before it is defined. *) let mkCaseEq a : unit Proofview.tactic = Proofview.Goal.nf_enter { enter = begin fun gl -> - let type_of_a = Tacmach.New.of_old (fun g -> Tacmach.pf_unsafe_type_of g a) gl in + let type_of_a = Tacmach.New.of_old (fun g -> Tacmach.pf_unsafe_type_of g (EConstr.of_constr a)) gl in Tacticals.New.tclTHENLIST [Tactics.generalize [mkApp(delayed_force refl_equal, [| type_of_a; a|])]; Proofview.Goal.nf_enter { enter = begin fun gl -> @@ -789,7 +789,7 @@ let destauto t = let destauto_in id = Proofview.Goal.nf_enter { enter = begin fun gl -> - let ctype = Tacmach.New.of_old (fun g -> Tacmach.pf_unsafe_type_of g (mkVar id)) gl in + let ctype = Tacmach.New.of_old (fun g -> Tacmach.pf_unsafe_type_of g (EConstr.mkVar id)) gl in (* Pp.msgnl (Printer.pr_lconstr (mkVar id)); *) (* Pp.msgnl (Printer.pr_lconstr (ctype)); *) destauto ctype diff --git a/ltac/rewrite.ml b/ltac/rewrite.ml index dfcfbfbd3..37b9a9edb 100644 --- a/ltac/rewrite.ml +++ b/ltac/rewrite.ml @@ -2188,7 +2188,7 @@ let setoid_transitivity c = let setoid_symmetry_in id = Proofview.V82.tactic (fun gl -> - let ctype = pf_unsafe_type_of gl (mkVar id) in + let ctype = pf_unsafe_type_of gl (EConstr.mkVar id) in let binders,concl = decompose_prod_assum ctype in let (equiv, args) = decompose_app concl in let rec split_last_two = function diff --git a/plugins/btauto/refl_btauto.ml b/plugins/btauto/refl_btauto.ml index 3ba5da149..1e49d8cad 100644 --- a/plugins/btauto/refl_btauto.ml +++ b/plugins/btauto/refl_btauto.ml @@ -179,6 +179,7 @@ module Btauto = struct let print_counterexample p env gl = let var = lapp witness [|p|] in + let var = EConstr.of_constr var in (* Compute an assignment that dissatisfies the goal *) let _, var = Tacmach.pf_reduction_of_red_expr gl (Genredexpr.CbvVm None) var in let rec to_list l = match decomp_term (Tacmach.project gl) l with diff --git a/plugins/cc/ccalgo.ml b/plugins/cc/ccalgo.ml index bc53b113d..102efe55b 100644 --- a/plugins/cc/ccalgo.ml +++ b/plugins/cc/ccalgo.ml @@ -508,7 +508,7 @@ let rec add_term state t= Not_found -> let b=next uf in let trm = constr_of_term t in - let typ = pf_unsafe_type_of state.gls trm in + let typ = pf_unsafe_type_of state.gls (EConstr.of_constr trm) in let typ = canonize_name typ in let new_node= match t with @@ -832,7 +832,7 @@ let complete_one_class state i= let id = new_state_var etyp state in app (Appli(t,Eps id)) (substl [mkVar id] rest) (n-1) in let _c = pf_unsafe_type_of state.gls - (constr_of_term (term state.uf pac.cnode)) in + (EConstr.of_constr (constr_of_term (term state.uf pac.cnode))) in let _args = List.map (fun i -> constr_of_term (term state.uf i)) pac.args in diff --git a/plugins/cc/cctac.ml b/plugins/cc/cctac.ml index 11da923e1..7c78f3a17 100644 --- a/plugins/cc/cctac.ml +++ b/plugins/cc/cctac.ml @@ -243,7 +243,8 @@ let app_global f args k = let new_app_global f args k = Tacticals.New.pf_constr_of_global (Lazy.force f) (fun fc -> k (mkApp (fc, args))) -let new_refine c = Proofview.V82.tactic (refine c) +let new_refine c = Proofview.V82.tactic (refine (EConstr.of_constr c)) +let refine c = refine (EConstr.of_constr c) let assert_before n c = Proofview.Goal.enter { enter = begin fun gl -> @@ -265,7 +266,7 @@ let refresh_universes ty k = let rec proof_tac p : unit Proofview.tactic = Proofview.Goal.nf_enter { enter = begin fun gl -> - let type_of t = Tacmach.New.pf_unsafe_type_of gl t in + let type_of t = Tacmach.New.pf_unsafe_type_of gl (EConstr.of_constr t) in try (* type_of can raise exceptions *) match p.p_rule with Ax c -> exact_check c @@ -336,7 +337,7 @@ let refute_tac c t1 t2 p = let neweq= new_app_global _eq [|intype;tt1;tt2|] in Tacticals.New.tclTHENS (neweq (assert_before (Name hid))) [proof_tac p; simplest_elim false_t] - in refresh_universes (Tacmach.New.pf_unsafe_type_of gl tt1) k + in refresh_universes (Tacmach.New.pf_unsafe_type_of gl (EConstr.of_constr tt1)) k end } let refine_exact_check c gl = @@ -354,7 +355,7 @@ let convert_to_goal_tac c t1 t2 p = let endt=app_global _eq_rect [|sort;tt1;identity;c;tt2;mkVar e|] in Tacticals.New.tclTHENS (neweq (assert_before (Name e))) [proof_tac p; Proofview.V82.tactic (endt refine_exact_check)] - in refresh_universes (Tacmach.New.pf_unsafe_type_of gl tt2) k + in refresh_universes (Tacmach.New.pf_unsafe_type_of gl (EConstr.of_constr tt2)) k end } let convert_to_hyp_tac c1 t1 c2 t2 p = @@ -376,7 +377,7 @@ let discriminate_tac (cstr,u as cstru) p = let identity = Universes.constr_of_global (Lazy.force _I) in let trivial = Universes.constr_of_global (Lazy.force _True) in let evm = Tacmach.New.project gl in - let evm, intype = refresh_type env evm (EConstr.of_constr (Tacmach.New.pf_unsafe_type_of gl t1)) in + let evm, intype = refresh_type env evm (EConstr.of_constr (Tacmach.New.pf_unsafe_type_of gl (EConstr.of_constr t1))) in let evm, outtype = Evd.new_sort_variable Evd.univ_flexible evm in let outtype = mkSort outtype in let pred = mkLambda(Name xid,outtype,mkRel 1) in diff --git a/plugins/decl_mode/decl_proof_instr.ml b/plugins/decl_mode/decl_proof_instr.ml index 3bb6f1b5d..031a6253a 100644 --- a/plugins/decl_mode/decl_proof_instr.ml +++ b/plugins/decl_mode/decl_proof_instr.ml @@ -491,7 +491,7 @@ let thus_tac c ctyp submetas gls = Proofview.V82.of_tactic (exact_check proof) gls else let refiner = concl_refiner list proof gls in - Tacmach.refine refiner gls + Tacmach.refine (EConstr.of_constr refiner) gls (* general forward step *) @@ -799,7 +799,7 @@ let rec take_tac wits gls = match wits with [] -> tclIDTAC gls | wit::rest -> - let typ = pf_unsafe_type_of gls wit in + let typ = pf_unsafe_type_of gls (EConstr.of_constr wit) in tclTHEN (thus_tac wit typ []) (take_tac rest) gls @@ -879,7 +879,7 @@ let start_tree env ind rp = let build_per_info etype casee gls = let concl=pf_concl gls in let env=pf_env gls in - let ctyp=pf_unsafe_type_of gls casee in + let ctyp=pf_unsafe_type_of gls (EConstr.of_constr casee) in let is_dep = dependent (project gls) (EConstr.of_constr casee) (EConstr.of_constr concl) in let hd,args = decompose_app (special_whd gls ctyp) in let (ind,u) = @@ -894,7 +894,7 @@ let build_per_info etype casee gls = | _ -> mind.mind_nparams,None in let params,real_args = List.chop nparams args in let abstract_obj c body = - let typ=pf_unsafe_type_of gls c in + let typ=pf_unsafe_type_of gls (EConstr.of_constr c) in lambda_create env (typ,subst_term (project gls) (EConstr.of_constr c) (EConstr.of_constr body)) in let pred= List.fold_right abstract_obj real_args (lambda_create env (ctyp,subst_term (project gls) (EConstr.of_constr casee) (EConstr.of_constr concl))) in @@ -1256,13 +1256,13 @@ let rec execute_cases fix_name per_info tacnext args objs nhrec tree gls = let nparams = mind.mind_nparams in let concl=pf_concl gls in let env=pf_env gls in - let ctyp=pf_unsafe_type_of gls casee in + let ctyp=pf_unsafe_type_of gls (EConstr.of_constr casee) in let hd,all_args = decompose_app (special_whd gls ctyp) in let ind', u = destInd hd in let _ = assert (eq_ind ind' ind) in (* just in case *) let params,real_args = List.chop nparams all_args in let abstract_obj c body = - let typ=pf_unsafe_type_of gls c in + let typ=pf_unsafe_type_of gls (EConstr.of_constr c) in lambda_create env (typ,subst_term (project gls) (EConstr.of_constr c) (EConstr.of_constr body)) in let elim_pred = List.fold_right abstract_obj real_args (lambda_create env (ctyp,subst_term (project gls) (EConstr.of_constr casee) (EConstr.of_constr concl))) in @@ -1314,7 +1314,7 @@ let rec execute_cases fix_name per_info tacnext args objs nhrec tree gls = execute_cases fix_name per_info tacnext p_args objs nhrec tree] gls0 in tclTHENSV - (refine case_term) + (refine (EConstr.of_constr case_term)) (Array.mapi branch_tac br) gls | Split_patt (_, _, _) , [] -> anomaly ~label:"execute_cases " (Pp.str "Nothing to split") diff --git a/plugins/firstorder/instances.ml b/plugins/firstorder/instances.ml index 44bdb585a..6c245063c 100644 --- a/plugins/firstorder/instances.ml +++ b/plugins/firstorder/instances.ml @@ -104,7 +104,7 @@ let mk_open_instance id idc gl m t= let evmap=Refiner.project gl in let var_id= if id==dummy_id then dummy_bvid else - let typ=pf_unsafe_type_of gl idc in + let typ=pf_unsafe_type_of gl (EConstr.of_constr idc) in (* since we know we will get a product, reduction is not too expensive *) let (nam,_,_)=destProd (whd_all env evmap (EConstr.of_constr typ)) in diff --git a/plugins/firstorder/sequent.ml b/plugins/firstorder/sequent.ml index 1248b60a7..87e7192d7 100644 --- a/plugins/firstorder/sequent.ml +++ b/plugins/firstorder/sequent.ml @@ -200,7 +200,7 @@ let extend_with_ref_list l seq gl = let l = expand_constructor_hints l in let f gr (seq,gl) = let gl, c = pf_eapply Evd.fresh_global gl gr in - let typ=(pf_unsafe_type_of gl c) in + let typ=(pf_unsafe_type_of gl (EConstr.of_constr c)) in (add_formula Hyp gr typ seq gl,gl) in List.fold_right f l (seq,gl) @@ -215,7 +215,7 @@ let extend_with_auto_hints l seq gl= let (c, _, _) = c in (try let gr = global_of_constr c in - let typ=(pf_unsafe_type_of gl c) in + let typ=(pf_unsafe_type_of gl (EConstr.of_constr c)) in seqref:=add_formula Hint gr typ !seqref gl with Not_found->()) | _-> () in diff --git a/plugins/funind/functional_principles_proofs.ml b/plugins/funind/functional_principles_proofs.ml index 83fc48623..b674f40e9 100644 --- a/plugins/funind/functional_principles_proofs.ml +++ b/plugins/funind/functional_principles_proofs.ml @@ -202,6 +202,7 @@ let prove_trivial_eq h_id context (constructor,type_of_term,term) = (List.map mkVar context_hyps) in let to_refine = applist(mkVar h_id,List.rev context_hyps') in + let to_refine = EConstr.of_constr to_refine in refine to_refine g ) ] @@ -329,7 +330,8 @@ let change_eq env sigma hyp_id (context:Context.Rel.t) x t end_of_type = let all_ids = pf_ids_of_hyps g in let new_ids,_ = list_chop ctxt_size all_ids in let to_refine = applist(witness_fun,List.rev_map mkVar new_ids) in - let evm, _ = pf_apply Typing.type_of g (EConstr.of_constr to_refine) in + let to_refine = EConstr.of_constr to_refine in + let evm, _ = pf_apply Typing.type_of g to_refine in tclTHEN (Refiner.tclEVARS evm) (refine to_refine) g ) in @@ -448,6 +450,7 @@ let clean_hyp_with_heq ptes_infos eq_hyps hyp_id env sigma = List.rev_map mkVar (rec_pte_id::context_hyps_ids) ) in + let to_refine = EConstr.of_constr to_refine in (* observe_tac "rec hyp " *) (tclTHENS (Proofview.V82.of_tactic (assert_before (Name rec_pte_id) t_x)) @@ -497,6 +500,7 @@ let clean_hyp_with_heq ptes_infos eq_hyps hyp_id env sigma = List.rev (coq_I::List.map mkVar context_hyps) ) in + let to_refine = (EConstr.of_constr to_refine) in refine to_refine g ) ] @@ -594,7 +598,7 @@ let treat_new_case ptes_infos nb_prod continue_tac term dyn_infos = tclMAP (fun id -> Proofview.V82.of_tactic (introduction ~check:false id)) dyn_infos.rec_hyps; observe_tac "after_introduction" (fun g' -> (* We get infos on the equations introduced*) - let new_term_value_eq = pf_unsafe_type_of g' (mkVar heq_id) in + let new_term_value_eq = pf_unsafe_type_of g' (EConstr.mkVar heq_id) in (* compute the new value of the body *) let new_term_value = match kind_of_term new_term_value_eq with @@ -605,13 +609,14 @@ let treat_new_case ptes_infos nb_prod continue_tac term dyn_infos = ); anomaly (Pp.str "cannot compute new term value") in + let term = EConstr.of_constr term in let fun_body = mkLambda(Anonymous, pf_unsafe_type_of g' term, - Termops.replace_term (project g') (EConstr.of_constr term) (EConstr.mkRel 1) (EConstr.of_constr dyn_infos.info) + Termops.replace_term (project g') term (EConstr.mkRel 1) (EConstr.of_constr dyn_infos.info) ) in - let new_body = pf_nf_betaiota g' (mkApp(fun_body,[| new_term_value |])) in + let new_body = pf_nf_betaiota g' (EConstr.of_constr (mkApp(fun_body,[| new_term_value |]))) in let new_infos = {dyn_infos with info = new_body; @@ -700,7 +705,7 @@ let build_proof let dyn_infos = {dyn_info' with info = mkCase(ci,ct,t,cb)} in let g_nb_prod = nb_prod (project g) (EConstr.of_constr (pf_concl g)) in - let type_of_term = pf_unsafe_type_of g t in + let type_of_term = pf_unsafe_type_of g (EConstr.of_constr t) in let term_eq = make_refl_eq (Lazy.force refl_equal) type_of_term t in @@ -741,7 +746,7 @@ let build_proof let id = pf_last_hyp g' |> get_id in let new_term = pf_nf_betaiota g' - (mkApp(dyn_infos.info,[|mkVar id|])) + (EConstr.of_constr (mkApp(dyn_infos.info,[|mkVar id|]))) in let new_infos = {dyn_infos with info = new_term} in let do_prove new_hyps = @@ -908,6 +913,7 @@ let prove_rec_hyp_for_struct fix_info = let rec_hyp_proof = mkApp(mkVar fix_info.name,array_get_start pte_args) in + let rec_hyp_proof = EConstr.of_constr rec_hyp_proof in refine rec_hyp_proof g )) @@ -921,7 +927,7 @@ let generalize_non_dep hyp g = (* observe (str "rec id := " ++ Ppconstr.pr_id hyp); *) let hyps = [hyp] in let env = Global.env () in - let hyp_typ = pf_unsafe_type_of g (mkVar hyp) in + let hyp_typ = pf_unsafe_type_of g (EConstr.mkVar hyp) in let to_revert,_ = let open Context.Named.Declaration in Environ.fold_named_context_reverse (fun (clear,keep) decl -> @@ -1418,7 +1424,7 @@ let backtrack_eqs_until_hrec hrec eqs : tactic = let rewrite = tclFIRST (List.map (fun x -> Proofview.V82.of_tactic (Equality.rewriteRL x)) eqs ) in - let _,hrec_concl = decompose_prod (pf_unsafe_type_of gls (mkVar hrec)) in + let _,hrec_concl = decompose_prod (pf_unsafe_type_of gls (EConstr.mkVar hrec)) in let f_app = Array.last (snd (destApp hrec_concl)) in let f = (fst (destApp f_app)) in let rec backtrack : tactic = diff --git a/plugins/funind/indfun.ml b/plugins/funind/indfun.ml index 0743fc5d9..e3ba52246 100644 --- a/plugins/funind/indfun.ml +++ b/plugins/funind/indfun.ml @@ -75,11 +75,11 @@ let functional_induction with_clean c princl pat = user_err (str "Cannot find induction principle for " ++Printer.pr_lconstr (mkConst c') ) in - (princ,NoBindings, Tacmach.pf_unsafe_type_of g' princ,g') + (princ,NoBindings, Tacmach.pf_unsafe_type_of g' (EConstr.of_constr princ),g') | _ -> raise (UserError(None,str "functional induction must be used with a function" )) end | Some ((princ,binding)) -> - princ,binding,Tacmach.pf_unsafe_type_of g princ,g + princ,binding,Tacmach.pf_unsafe_type_of g (EConstr.of_constr princ),g in let princ_infos = Tactics.compute_elim_sig princ_type in let args_as_induction_constr = diff --git a/plugins/funind/invfun.ml b/plugins/funind/invfun.ml index e5286fb1f..635925562 100644 --- a/plugins/funind/invfun.ml +++ b/plugins/funind/invfun.ml @@ -305,7 +305,7 @@ let prove_fun_correct evd functional_induction funs_constr graphs_constr schemes let constructor_args g = List.fold_right (fun hid acc -> - let type_of_hid = pf_unsafe_type_of g (mkVar hid) in + let type_of_hid = pf_unsafe_type_of g (EConstr.mkVar hid) in match kind_of_term type_of_hid with | Prod(_,_,t') -> begin @@ -596,7 +596,7 @@ let rec reflexivity_with_destruct_cases g = match sc with None -> tclIDTAC g | Some id -> - match kind_of_term (pf_unsafe_type_of g (mkVar id)) with + match kind_of_term (pf_unsafe_type_of g (EConstr.mkVar id)) with | App(eq,[|_;t1;t2|]) when eq_constr eq eq_ind -> if Equality.discriminable (pf_env g) (project g) t1 t2 then Proofview.V82.of_tactic (Equality.discrHyp id) g @@ -661,7 +661,7 @@ let prove_fun_complete funcs graphs schemes lemmas_types_infos i : tactic = (* We get the constant and the principle corresponding to this lemma *) let f = funcs.(i) in let graph_principle = nf_zeta (EConstr.of_constr schemes.(i)) in - let princ_type = pf_unsafe_type_of g graph_principle in + let princ_type = pf_unsafe_type_of g (EConstr.of_constr graph_principle) in let princ_infos = Tactics.compute_elim_sig princ_type in (* Then we get the number of argument of the function and compute a fresh name for each of them @@ -919,7 +919,7 @@ let derive_correctness make_scheme functional_induction (funs: pconstant list) ( if the type of hypothesis has not this form or if we cannot find the completeness lemma then we do nothing *) let revert_graph kn post_tac hid g = - let typ = pf_unsafe_type_of g (mkVar hid) in + let typ = pf_unsafe_type_of g (EConstr.mkVar hid) in match kind_of_term typ with | App(i,args) when isInd i -> let ((kn',num) as ind'),u = destInd i in @@ -970,7 +970,7 @@ let revert_graph kn post_tac hid g = let functional_inversion kn hid fconst f_correct : tactic = fun g -> let old_ids = List.fold_right Id.Set.add (pf_ids_of_hyps g) Id.Set.empty in - let type_of_h = pf_unsafe_type_of g (mkVar hid) in + let type_of_h = pf_unsafe_type_of g (EConstr.mkVar hid) in match kind_of_term type_of_h with | App(eq,args) when eq_constr eq (make_eq ()) -> let pre_tac,f_args,res = @@ -1022,7 +1022,7 @@ let invfun qhyp f g = Proofview.V82.of_tactic begin Tactics.try_intros_until (fun hid -> Proofview.V82.tactic begin fun g -> - let hyp_typ = pf_unsafe_type_of g (mkVar hid) in + let hyp_typ = pf_unsafe_type_of g (EConstr.mkVar hid) in match kind_of_term hyp_typ with | App(eq,args) when eq_constr eq (make_eq ()) -> begin diff --git a/plugins/funind/recdef.ml b/plugins/funind/recdef.ml index bdbf0242d..b2c93a754 100644 --- a/plugins/funind/recdef.ml +++ b/plugins/funind/recdef.ml @@ -117,7 +117,7 @@ let pf_get_new_ids idl g = let compute_renamed_type gls c = rename_bound_vars_as_displayed (*no avoid*) [] (*no rels*) [] - (pf_unsafe_type_of gls c) + (pf_unsafe_type_of gls (EConstr.of_constr c)) let h'_id = Id.of_string "h'" let teq_id = Id.of_string "teq" let ano_id = Id.of_string "anonymous" @@ -402,7 +402,7 @@ let treat_case forbid_new_ids to_intros finalize_tac nb_lam e infos : tactic = Proofview.V82.of_tactic (clear to_intros); h_intros to_intros; (fun g' -> - let ty_teq = pf_unsafe_type_of g' (mkVar heq) in + let ty_teq = pf_unsafe_type_of g' (EConstr.mkVar heq) in let teq_lhs,teq_rhs = let _,args = try destApp ty_teq with DestKO -> assert false in args.(1),args.(2) @@ -516,13 +516,13 @@ let rec prove_lt hyple g = in let h = List.find (fun id -> - match decompose_app (pf_unsafe_type_of g (mkVar id)) with + match decompose_app (pf_unsafe_type_of g (EConstr.mkVar id)) with | _, t::_ -> eq_constr t varx | _ -> false ) hyple in let y = - List.hd (List.tl (snd (decompose_app (pf_unsafe_type_of g (mkVar h))))) in + List.hd (List.tl (snd (decompose_app (pf_unsafe_type_of g (EConstr.mkVar h))))) in observe_tclTHENLIST (str "prove_lt1")[ Proofview.V82.of_tactic (apply (mkApp(le_lt_trans (),[|varx;y;varz;mkVar h|]))); observe_tac (str "prove_lt") (prove_lt hyple) @@ -684,7 +684,7 @@ let mkDestructEq : if Id.List.mem id not_on_hyp || not (Termops.occur_term (project g) (EConstr.of_constr expr) (EConstr.of_constr (get_type decl))) then None else Some id) hyps in let to_revert_constr = List.rev_map mkVar to_revert in - let type_of_expr = pf_unsafe_type_of g expr in + let type_of_expr = pf_unsafe_type_of g (EConstr.of_constr expr) in let new_hyps = mkApp(Lazy.force refl_equal, [|type_of_expr; expr|]):: to_revert_constr in pf_typel new_hyps (fun _ -> @@ -839,7 +839,7 @@ let rec prove_le g = let matching_fun = pf_is_matching g (Pattern.PApp(Pattern.PRef (reference_of_constr (le ())),[|Pattern.PVar (destVar x);Pattern.PMeta None|])) in - let (h,t) = List.find (fun (_,t) -> matching_fun t) (pf_hyps_types g) + let (h,t) = List.find (fun (_,t) -> matching_fun (EConstr.of_constr t)) (pf_hyps_types g) in let y = let _,args = decompose_app t in diff --git a/plugins/omega/coq_omega.ml b/plugins/omega/coq_omega.ml index d15449aef..b832250a5 100644 --- a/plugins/omega/coq_omega.ml +++ b/plugins/omega/coq_omega.ml @@ -568,7 +568,7 @@ let abstract_path typ path t = mkLambda (Name (Id.of_string "x"), typ, abstract), !term_occur let focused_simpl path gl = - let newc = context (fun i t -> pf_nf gl t) (List.rev path) (pf_concl gl) in + let newc = context (fun i t -> pf_nf gl (EConstr.of_constr t)) (List.rev path) (pf_concl gl) in Proofview.V82.of_tactic (convert_concl_no_check newc DEFAULTcast) gl let focused_simpl path = focused_simpl path @@ -644,7 +644,7 @@ let clever_rewrite_base_poly typ p result theorem gl = [| typ; result; mkRel 2; mkRel 1; occ; theorem |]))), [abstracted]) in - exact (applist(t,[mkNewMeta()])) gl + exact (EConstr.of_constr (applist(t,[mkNewMeta()]))) gl let clever_rewrite_base p result theorem gl = clever_rewrite_base_poly (Lazy.force coq_Z) p result theorem gl @@ -665,7 +665,7 @@ let clever_rewrite p vpath t gl = let (abstracted,occ) = abstract_path (Lazy.force coq_Z) (List.rev p) full in let vargs = List.map (fun p -> occurrence p occ) vpath in let t' = applist(t, (vargs @ [abstracted])) in - exact (applist(t',[mkNewMeta()])) gl + exact (EConstr.of_constr (applist(t',[mkNewMeta()]))) gl let rec shuffle p (t1,t2) = match t1,t2 with @@ -1384,7 +1384,7 @@ let destructure_omega gl tac_def (id,c) = else try match destructurate_prop c with | Kapp(Eq,[typ;t1;t2]) - when begin match destructurate_type (pf_nf gl typ) with Kapp(Z,[]) -> true | _ -> false end -> + when begin match destructurate_type (pf_nf gl (EConstr.of_constr typ)) with Kapp(Z,[]) -> true | _ -> false end -> let t = mk_plus t1 (mk_inv t2) in normalize_equation id EQUA (Lazy.force coq_Zegal_left) 2 t t1 t2 tac_def @@ -1659,7 +1659,7 @@ let rec decidability gl t = | Kapp(Not,[t1]) -> mkApp (Lazy.force coq_dec_not, [| t1; decidability gl t1 |]) | Kapp(Eq,[typ;t1;t2]) -> - begin match destructurate_type (pf_nf gl typ) with + begin match destructurate_type (pf_nf gl (EConstr.of_constr typ)) with | Kapp(Z,[]) -> mkApp (Lazy.force coq_dec_eq, [| t1;t2 |]) | Kapp(Nat,[]) -> mkApp (Lazy.force coq_dec_eq_nat, [| t1;t2 |]) | _ -> raise Undecidable @@ -1720,7 +1720,7 @@ let destructure_hyps = | Kimp(t1,t2) -> (* t1 and t2 might be in Type rather than Prop. For t1, the decidability check will ensure being Prop. *) - if is_Prop (type_of t2) + if is_Prop (type_of (EConstr.of_constr t2)) then let d1 = decidability t1 in Tacticals.New.tclTHENLIST [ @@ -1789,7 +1789,7 @@ let destructure_hyps = with Not_found -> loop lit) | Kapp(Eq,[typ;t1;t2]) -> if !old_style_flag then begin - match destructurate_type (pf_nf typ) with + match destructurate_type (pf_nf (EConstr.of_constr typ)) with | Kapp(Nat,_) -> Tacticals.New.tclTHENLIST [ (simplest_elim @@ -1806,7 +1806,7 @@ let destructure_hyps = ] | _ -> loop lit end else begin - match destructurate_type (pf_nf typ) with + match destructurate_type (pf_nf (EConstr.of_constr typ)) with | Kapp(Nat,_) -> (Tacticals.New.tclTHEN (convert_hyp_no_check (NamedDecl.set_type (mkApp (Lazy.force coq_neq, [| t1;t2|])) @@ -1849,7 +1849,7 @@ let destructure_goal = let dec = decidability t in Tacticals.New.tclTHEN (Proofview.V82.tactic (Tacmach.refine - (mkApp (Lazy.force coq_dec_not_not, [| t; dec; mkNewMeta () |])))) + (EConstr.of_constr (mkApp (Lazy.force coq_dec_not_not, [| t; dec; mkNewMeta () |]))))) intro with Undecidable -> Tactics.elim_type (build_coq_False ()) in diff --git a/plugins/romega/const_omega.ml b/plugins/romega/const_omega.ml index 4935fe4bb..f2d91bad3 100644 --- a/plugins/romega/const_omega.ml +++ b/plugins/romega/const_omega.ml @@ -353,7 +353,7 @@ let parse_term t = let parse_rel gl t = try match destructurate t with | Kapp("eq",[typ;t1;t2]) - when destructurate (Tacmach.pf_nf gl typ) = Kapp("Z",[]) -> Req (t1,t2) + when destructurate (Tacmach.pf_nf gl (EConstr.of_constr typ)) = Kapp("Z",[]) -> Req (t1,t2) | Kapp("Zne",[t1;t2]) -> Rne (t1,t2) | Kapp("Z.le",[t1;t2]) -> Rle (t1,t2) | Kapp("Z.lt",[t1;t2]) -> Rlt (t1,t2) diff --git a/plugins/ssrmatching/ssrmatching.ml4 b/plugins/ssrmatching/ssrmatching.ml4 index 7f628f165..ace557a52 100644 --- a/plugins/ssrmatching/ssrmatching.ml4 +++ b/plugins/ssrmatching/ssrmatching.ml4 @@ -1390,7 +1390,7 @@ let ssrpatterntac _ist (arg_ist,arg) gl = let concl0 = pf_concl gl in let (t, uc), concl_x = fill_occ_pattern (Global.env()) sigma0 concl0 pat noindex 1 in - let gl, tty = pf_type_of gl t in + let gl, tty = pf_type_of gl (EConstr.of_constr t) in let concl = mkLetIn (Name (id_of_string "selected"), t, tty, concl_x) in Proofview.V82.of_tactic (convert_concl concl DEFAULTcast) gl diff --git a/pretyping/tacred.ml b/pretyping/tacred.ml index ae53c12ae..24d4af89a 100644 --- a/pretyping/tacred.ml +++ b/pretyping/tacred.ml @@ -1214,8 +1214,8 @@ let reduce_to_ind_gen allow_product env sigma t = in elimrec env t [] -let reduce_to_quantified_ind env sigma c = reduce_to_ind_gen true env sigma (EConstr.of_constr c) -let reduce_to_atomic_ind env sigma c = reduce_to_ind_gen false env sigma (EConstr.of_constr c) +let reduce_to_quantified_ind env sigma c = reduce_to_ind_gen true env sigma c +let reduce_to_atomic_ind env sigma c = reduce_to_ind_gen false env sigma c let find_hnf_rectype env sigma t = let ind,t = reduce_to_atomic_ind env sigma t in diff --git a/pretyping/tacred.mli b/pretyping/tacred.mli index d32fcf491..3587ae281 100644 --- a/pretyping/tacred.mli +++ b/pretyping/tacred.mli @@ -75,12 +75,12 @@ val cbv_norm_flags : CClosure.RedFlags.reds -> reduction_function (** [reduce_to_atomic_ind env sigma t] puts [t] in the form [t'=(I args)] with [I] an inductive definition; returns [I] and [t'] or fails with a user error *) -val reduce_to_atomic_ind : env -> evar_map -> types -> pinductive * types +val reduce_to_atomic_ind : env -> evar_map -> EConstr.types -> pinductive * types (** [reduce_to_quantified_ind env sigma t] puts [t] in the form [t'=(x1:A1)..(xn:An)(I args)] with [I] an inductive definition; returns [I] and [t'] or fails with a user error *) -val reduce_to_quantified_ind : env -> evar_map -> types -> pinductive * types +val reduce_to_quantified_ind : env -> evar_map -> EConstr.types -> pinductive * types (** [reduce_to_quantified_ref env sigma ref t] try to put [t] in the form [t'=(x1:A1)..(xn:An)(ref args)] and fails with user error if not possible *) @@ -91,7 +91,7 @@ val reduce_to_atomic_ref : env -> evar_map -> global_reference -> EConstr.types -> types val find_hnf_rectype : - env -> evar_map -> types -> pinductive * constr list + env -> evar_map -> EConstr.types -> pinductive * constr list val contextually : bool -> occurrences * constr_pattern -> (patvar_map -> reduction_function) -> reduction_function diff --git a/proofs/clenvtac.ml b/proofs/clenvtac.ml index bc5f17bf5..5b9322bfe 100644 --- a/proofs/clenvtac.ml +++ b/proofs/clenvtac.ml @@ -95,7 +95,7 @@ let clenv_refine with_evars ?(with_classes=true) clenv = let clenv = { clenv with evd = evd' } in tclTHEN (tclEVARS (Evd.clear_metas evd')) - (refine_no_check (clenv_cast_meta clenv (clenv_value clenv))) gl + (refine_no_check (EConstr.of_constr (clenv_cast_meta clenv (clenv_value clenv)))) gl end open Unification diff --git a/proofs/logic.ml b/proofs/logic.ml index c5d6e3e08..829c96296 100644 --- a/proofs/logic.ml +++ b/proofs/logic.ml @@ -490,6 +490,7 @@ and mk_arggoals sigma goal goalacc funty allargs = and mk_casegoals sigma goal goalacc p c = let env = Goal.V82.env sigma goal in let (acc',ct,sigma,c') = mk_hdgoals sigma goal goalacc c in + let ct = EConstr.of_constr ct in let (acc'',pt,sigma,p') = mk_hdgoals sigma goal acc' p in let indspec = try Tacred.find_hnf_rectype env sigma ct diff --git a/proofs/tacmach.ml b/proofs/tacmach.ml index c2ebb76d7..b732e5b9d 100644 --- a/proofs/tacmach.ml +++ b/proofs/tacmach.ml @@ -77,37 +77,36 @@ let pf_global gls id = Constrintern.construct_reference (pf_hyps gls) id let pf_reduction_of_red_expr gls re c = let (redfun, _) = reduction_of_red_expr (pf_env gls) re in let sigma = Sigma.Unsafe.of_evar_map (project gls) in - let Sigma (c, sigma, _) = redfun.e_redfun (pf_env gls) sigma (EConstr.of_constr c) in + let Sigma (c, sigma, _) = redfun.e_redfun (pf_env gls) sigma c in (Sigma.to_evar_map sigma, c) let pf_apply f gls = f (pf_env gls) (project gls) let pf_eapply f gls x = on_sig gls (fun evm -> f (pf_env gls) evm x) let pf_reduce = pf_apply -let pf_reduce' f gl c = pf_apply f gl (EConstr.of_constr c) let pf_e_reduce = pf_apply -let pf_whd_all = pf_reduce' whd_all -let pf_hnf_constr = pf_reduce' hnf_constr -let pf_nf = pf_reduce' simpl -let pf_nf_betaiota = pf_reduce' (fun _ -> nf_betaiota) -let pf_compute = pf_reduce' compute -let pf_unfoldn ubinds = pf_reduce' (unfoldn ubinds) -let pf_unsafe_type_of = pf_reduce' unsafe_type_of -let pf_type_of = pf_reduce' type_of +let pf_whd_all = pf_reduce whd_all +let pf_hnf_constr = pf_reduce hnf_constr +let pf_nf = pf_reduce simpl +let pf_nf_betaiota = pf_reduce (fun _ -> nf_betaiota) +let pf_compute = pf_reduce compute +let pf_unfoldn ubinds = pf_reduce (unfoldn ubinds) +let pf_unsafe_type_of = pf_reduce unsafe_type_of +let pf_type_of = pf_reduce type_of let pf_get_type_of = pf_reduce Retyping.get_type_of -let pf_conv_x gl = pf_apply test_conversion gl Reduction.CONV +let pf_conv_x gl = pf_reduce test_conversion gl Reduction.CONV let pf_conv_x_leq gl = pf_reduce test_conversion gl Reduction.CUMUL let pf_const_value = pf_reduce (fun env _ -> constant_value_in env) let pf_reduce_to_quantified_ind = pf_reduce reduce_to_quantified_ind let pf_reduce_to_atomic_ind = pf_reduce reduce_to_atomic_ind -let pf_hnf_type_of gls = EConstr.of_constr %> pf_get_type_of gls %> pf_whd_all gls +let pf_hnf_type_of gls = pf_get_type_of gls %> EConstr.of_constr %> pf_whd_all gls -let pf_is_matching gl p c = pf_apply Constr_matching.is_matching_conv gl p (EConstr.of_constr c) -let pf_matches gl p c = pf_apply Constr_matching.matches_conv gl p (EConstr.of_constr c) +let pf_is_matching gl p c = pf_apply Constr_matching.is_matching_conv gl p c +let pf_matches gl p c = pf_apply Constr_matching.matches_conv gl p c (********************************************) (* Definition of the most primitive tactics *) @@ -116,12 +115,15 @@ let pf_matches gl p c = pf_apply Constr_matching.matches_conv gl p (EC let refiner = refiner let internal_cut_no_check replace id t gl = + let t = EConstr.Unsafe.to_constr t in refiner (Cut (true,replace,id,t)) gl let internal_cut_rev_no_check replace id t gl = + let t = EConstr.Unsafe.to_constr t in refiner (Cut (false,replace,id,t)) gl let refine_no_check c gl = + let c = EConstr.Unsafe.to_constr c in refiner (Refine c) gl (* Versions with consistency checks *) @@ -159,9 +161,6 @@ module New = struct let pf_apply f gl = f (Proofview.Goal.env gl) (project gl) - let pf_reduce f gl c = - f (Proofview.Goal.env gl) (project gl) (EConstr.of_constr c) - let of_old f gl = f { Evd.it = Proofview.Goal.goal gl ; sigma = project gl; } @@ -175,10 +174,10 @@ module New = struct let pf_concl = Proofview.Goal.concl let pf_unsafe_type_of gl t = - pf_apply unsafe_type_of gl (EConstr.of_constr t) + pf_apply unsafe_type_of gl t let pf_type_of gl t = - pf_apply type_of gl (EConstr.of_constr t) + pf_apply type_of gl t let pf_conv_x gl t1 t2 = pf_apply is_conv gl t1 t2 @@ -221,18 +220,18 @@ module New = struct let sigma = project gl in nf_evar sigma concl - let pf_whd_all gl t = pf_reduce whd_all gl t + let pf_whd_all gl t = pf_apply whd_all gl t let pf_get_type_of gl t = pf_apply Retyping.get_type_of gl t let pf_reduce_to_quantified_ind gl t = pf_apply reduce_to_quantified_ind gl t - let pf_hnf_constr gl t = pf_reduce hnf_constr gl t + let pf_hnf_constr gl t = pf_apply hnf_constr gl t let pf_hnf_type_of gl t = - pf_whd_all gl (pf_get_type_of gl (EConstr.of_constr t)) + pf_whd_all gl (EConstr.of_constr (pf_get_type_of gl t)) - let pf_matches gl pat t = pf_apply Constr_matching.matches_conv gl pat (EConstr.of_constr t) + let pf_matches gl pat t = pf_apply Constr_matching.matches_conv gl pat t let pf_whd_all gl t = pf_apply whd_all gl t let pf_compute gl t = pf_apply compute gl t diff --git a/proofs/tacmach.mli b/proofs/tacmach.mli index 16f6f88c1..07d02212c 100644 --- a/proofs/tacmach.mli +++ b/proofs/tacmach.mli @@ -40,9 +40,9 @@ val pf_nth_hyp_id : goal sigma -> int -> Id.t val pf_last_hyp : goal sigma -> Context.Named.Declaration.t val pf_ids_of_hyps : goal sigma -> Id.t list val pf_global : goal sigma -> Id.t -> constr -val pf_unsafe_type_of : goal sigma -> constr -> types -val pf_type_of : goal sigma -> constr -> evar_map * types -val pf_hnf_type_of : goal sigma -> constr -> types +val pf_unsafe_type_of : goal sigma -> EConstr.constr -> types +val pf_type_of : goal sigma -> EConstr.constr -> evar_map * types +val pf_hnf_type_of : goal sigma -> EConstr.constr -> types val pf_get_hyp : goal sigma -> Id.t -> Context.Named.Declaration.t val pf_get_hyp_typ : goal sigma -> Id.t -> types @@ -50,7 +50,7 @@ val pf_get_hyp_typ : goal sigma -> Id.t -> types val pf_get_new_id : Id.t -> goal sigma -> Id.t val pf_get_new_ids : Id.t list -> goal sigma -> Id.t list -val pf_reduction_of_red_expr : goal sigma -> red_expr -> constr -> evar_map * constr +val pf_reduction_of_red_expr : goal sigma -> red_expr -> EConstr.constr -> evar_map * constr val pf_apply : (env -> evar_map -> 'a) -> goal sigma -> 'a @@ -63,35 +63,35 @@ val pf_e_reduce : (env -> evar_map -> constr -> evar_map * constr) -> goal sigma -> constr -> evar_map * constr -val pf_whd_all : goal sigma -> constr -> constr -val pf_hnf_constr : goal sigma -> constr -> constr -val pf_nf : goal sigma -> constr -> constr -val pf_nf_betaiota : goal sigma -> constr -> constr -val pf_reduce_to_quantified_ind : goal sigma -> types -> pinductive * types -val pf_reduce_to_atomic_ind : goal sigma -> types -> pinductive * types -val pf_compute : goal sigma -> constr -> constr +val pf_whd_all : goal sigma -> EConstr.constr -> constr +val pf_hnf_constr : goal sigma -> EConstr.constr -> constr +val pf_nf : goal sigma -> EConstr.constr -> constr +val pf_nf_betaiota : goal sigma -> EConstr.constr -> constr +val pf_reduce_to_quantified_ind : goal sigma -> EConstr.types -> pinductive * types +val pf_reduce_to_atomic_ind : goal sigma -> EConstr.types -> pinductive * types +val pf_compute : goal sigma -> EConstr.constr -> constr val pf_unfoldn : (occurrences * evaluable_global_reference) list - -> goal sigma -> constr -> constr + -> goal sigma -> EConstr.constr -> constr val pf_const_value : goal sigma -> pconstant -> constr val pf_conv_x : goal sigma -> constr -> constr -> bool val pf_conv_x_leq : goal sigma -> constr -> constr -> bool -val pf_matches : goal sigma -> constr_pattern -> constr -> patvar_map -val pf_is_matching : goal sigma -> constr_pattern -> constr -> bool +val pf_matches : goal sigma -> constr_pattern -> EConstr.constr -> patvar_map +val pf_is_matching : goal sigma -> constr_pattern -> EConstr.constr -> bool (** {6 The most primitive tactics. } *) val refiner : rule -> tactic -val internal_cut_no_check : bool -> Id.t -> types -> tactic -val refine_no_check : constr -> tactic +val internal_cut_no_check : bool -> Id.t -> EConstr.types -> tactic +val refine_no_check : EConstr.constr -> tactic (** {6 The most primitive tactics with consistency and type checking } *) -val internal_cut : bool -> Id.t -> types -> tactic -val internal_cut_rev : bool -> Id.t -> types -> tactic -val refine : constr -> tactic +val internal_cut : bool -> Id.t -> EConstr.types -> tactic +val internal_cut_rev : bool -> Id.t -> EConstr.types -> tactic +val refine : EConstr.constr -> tactic (** {6 Pretty-printing functions (debug only). } *) val pr_gls : goal sigma -> Pp.std_ppcmds @@ -108,8 +108,8 @@ module New : sig val pf_env : ('a, 'r) Proofview.Goal.t -> Environ.env val pf_concl : ([ `NF ], 'r) Proofview.Goal.t -> types - val pf_unsafe_type_of : ('a, 'r) Proofview.Goal.t -> Term.constr -> Term.types - val pf_type_of : ('a, 'r) Proofview.Goal.t -> Term.constr -> evar_map * Term.types + val pf_unsafe_type_of : ('a, 'r) Proofview.Goal.t -> EConstr.constr -> Term.types + val pf_type_of : ('a, 'r) Proofview.Goal.t -> EConstr.constr -> evar_map * Term.types val pf_conv_x : ('a, 'r) Proofview.Goal.t -> EConstr.t -> EConstr.t -> bool val pf_get_new_id : identifier -> ([ `NF ], 'r) Proofview.Goal.t -> identifier @@ -121,15 +121,15 @@ module New : sig val pf_last_hyp : ([ `NF ], 'r) Proofview.Goal.t -> Context.Named.Declaration.t val pf_nf_concl : ([ `LZ ], 'r) Proofview.Goal.t -> types - val pf_reduce_to_quantified_ind : ('a, 'r) Proofview.Goal.t -> types -> pinductive * types + val pf_reduce_to_quantified_ind : ('a, 'r) Proofview.Goal.t -> EConstr.types -> pinductive * types - val pf_hnf_constr : ('a, 'r) Proofview.Goal.t -> constr -> types - val pf_hnf_type_of : ('a, 'r) Proofview.Goal.t -> constr -> types + val pf_hnf_constr : ('a, 'r) Proofview.Goal.t -> EConstr.constr -> types + val pf_hnf_type_of : ('a, 'r) Proofview.Goal.t -> EConstr.constr -> types val pf_whd_all : ('a, 'r) Proofview.Goal.t -> EConstr.t -> constr val pf_compute : ('a, 'r) Proofview.Goal.t -> EConstr.t -> constr - val pf_matches : ('a, 'r) Proofview.Goal.t -> constr_pattern -> constr -> patvar_map + val pf_matches : ('a, 'r) Proofview.Goal.t -> constr_pattern -> EConstr.constr -> patvar_map val pf_nf_evar : ('a, 'r) Proofview.Goal.t -> constr -> constr diff --git a/tactics/class_tactics.ml b/tactics/class_tactics.ml index b0f355170..a2699ba8d 100644 --- a/tactics/class_tactics.ml +++ b/tactics/class_tactics.ml @@ -226,7 +226,7 @@ let e_give_exact flags poly (c,clenv) gl = c, {gl with sigma = evd} else c, gl in - let t1 = pf_unsafe_type_of gl c in + let t1 = pf_unsafe_type_of gl (EConstr.of_constr c) in Proofview.V82.of_tactic (Clenvtac.unify ~flags t1 <*> exact_no_check c) gl let unify_e_resolve poly flags = { enter = begin fun gls (c,_,clenv) -> @@ -1514,7 +1514,7 @@ let is_ground c gl = let autoapply c i gl = let flags = auto_unif_flags Evar.Set.empty (Hints.Hint_db.transparent_state (Hints.searchtable_map i)) in - let cty = pf_unsafe_type_of gl c in + let cty = pf_unsafe_type_of gl (EConstr.of_constr c) in let ce = mk_clenv_from gl (c,cty) in let tac = { enter = fun gl -> (unify_e_resolve false flags).enter gl ((c,cty,Univ.ContextSet.empty),0,ce) } in diff --git a/tactics/contradiction.ml b/tactics/contradiction.ml index 9580fdbfc..2058b95a6 100644 --- a/tactics/contradiction.ml +++ b/tactics/contradiction.ml @@ -113,7 +113,7 @@ let contradiction_term (c,lbind as cl) = let sigma = Tacmach.New.project gl in let env = Proofview.Goal.env gl in let type_of = Tacmach.New.pf_unsafe_type_of gl in - let typ = type_of c in + let typ = type_of (EConstr.of_constr c) in let _, ccl = splay_prod env sigma (EConstr.of_constr typ) in if is_empty_type sigma ccl then Tacticals.New.tclTHEN diff --git a/tactics/eauto.ml b/tactics/eauto.ml index 0869ac0c7..2fad4fcf7 100644 --- a/tactics/eauto.ml +++ b/tactics/eauto.ml @@ -30,7 +30,7 @@ let eauto_unif_flags = auto_flags_of_state full_transparent_state let e_give_exact ?(flags=eauto_unif_flags) c = Proofview.Goal.enter { enter = begin fun gl -> - let t1 = Tacmach.New.pf_unsafe_type_of gl c in + let t1 = Tacmach.New.pf_unsafe_type_of gl (EConstr.of_constr c) in let t2 = Tacmach.New.pf_concl (Proofview.Goal.assume gl) in let sigma = Tacmach.New.project gl in if occur_existential sigma (EConstr.of_constr t1) || occur_existential sigma (EConstr.of_constr t2) then diff --git a/tactics/elim.ml b/tactics/elim.ml index b830ccefe..bcb1c05cc 100644 --- a/tactics/elim.ml +++ b/tactics/elim.ml @@ -80,7 +80,7 @@ let general_decompose recognizer c = Proofview.Goal.enter { enter = begin fun gl -> let type_of = pf_unsafe_type_of gl in let sigma = project gl in - let typc = type_of c in + let typc = type_of (EConstr.of_constr c) in tclTHENS (cut typc) [ tclTHEN (intro_using tmphyp_name) (onLastHypId @@ -133,7 +133,7 @@ let induction_trailer abs_i abs_j bargs = (onLastHypId (fun id -> Proofview.Goal.nf_enter { enter = begin fun gl -> - let idty = pf_unsafe_type_of gl (mkVar id) in + let idty = pf_unsafe_type_of gl (EConstr.mkVar id) in let fvty = global_vars (pf_env gl) (project gl) (EConstr.of_constr idty) in let possible_bring_hyps = (List.tl (nLastDecls gl (abs_j - abs_i))) @ bargs.Tacticals.assums diff --git a/tactics/eqdecide.ml b/tactics/eqdecide.ml index 1554d43f0..d1b14a907 100644 --- a/tactics/eqdecide.ml +++ b/tactics/eqdecide.ml @@ -156,7 +156,7 @@ let rec solveArg hyps eqonleft op largs rargs = match largs, rargs with ] | a1 :: largs, a2 :: rargs -> Proofview.Goal.enter { enter = begin fun gl -> - let rectype = pf_unsafe_type_of gl a1 in + let rectype = pf_unsafe_type_of gl (EConstr.of_constr a1) in let decide = mkDecideEqGoal eqonleft op rectype a1 a2 in let tac hyp = solveArg (hyp :: hyps) eqonleft op largs rargs in let subtacs = @@ -226,7 +226,7 @@ let decideEquality rectype = let compare c1 c2 = Proofview.Goal.enter { enter = begin fun gl -> - let rectype = pf_unsafe_type_of gl c1 in + let rectype = pf_unsafe_type_of gl (EConstr.of_constr c1) in let decide = mkDecideEqGoal true (build_coq_sumbool ()) rectype c1 c2 in (tclTHENS (cut decide) [(tclTHEN intro diff --git a/tactics/equality.ml b/tactics/equality.ml index 64b56b99b..ad80d2d1f 100644 --- a/tactics/equality.ml +++ b/tactics/equality.ml @@ -181,8 +181,8 @@ let instantiate_lemma_all frzevars gl c ty l l2r concl = in List.map try_occ occs let instantiate_lemma gl c ty l l2r concl = - let sigma, ct = pf_type_of gl c in - let t = try snd (reduce_to_quantified_ind (pf_env gl) sigma ct) with UserError _ -> ct in + let sigma, ct = pf_type_of gl (EConstr.of_constr c) in + let t = try snd (reduce_to_quantified_ind (pf_env gl) sigma (EConstr.of_constr ct)) with UserError _ -> ct in let eqclause = Clenv.make_clenv_binding (pf_env gl) sigma (c,t) l in [eqclause] @@ -992,6 +992,7 @@ let discr_positions env sigma (lbeq,eqn,(t,t1,t2)) eq_clause cpath dirn = let pf_ty = mkArrow eqn absurd_term in let absurd_clause = apply_on_clause (pf,pf_ty) eq_clause in let pf = Clenvtac.clenv_value_cast_meta absurd_clause in + let pf = EConstr.of_constr pf in Proofview.Unsafe.tclEVARS sigma <*> Proofview.tclEFFECTS eff <*> tclTHENS (assert_after Anonymous absurd_term) @@ -1012,8 +1013,8 @@ let onEquality with_evars tac (c,lbindc) = Proofview.Goal.nf_enter { enter = begin fun gl -> let type_of = pf_unsafe_type_of gl in let reduce_to_quantified_ind = pf_apply Tacred.reduce_to_quantified_ind gl in - let t = type_of c in - let t' = try snd (reduce_to_quantified_ind t) with UserError _ -> t in + let t = type_of (EConstr.of_constr c) in + let t' = try snd (reduce_to_quantified_ind (EConstr.of_constr t)) with UserError _ -> t in let eq_clause = pf_apply make_clenv_binding gl (c,t') lbindc in let eq_clause' = Clenvtac.clenv_pose_dependent_evars with_evars eq_clause in let eqn = clenv_type eq_clause' in @@ -1327,7 +1328,7 @@ let inject_if_homogenous_dependent_pair ty = if not (Ind_tables.check_scheme (!eq_dec_scheme_kind_name()) (fst ind) && pf_apply is_conv gl (EConstr.of_constr ar1.(2)) (EConstr.of_constr ar2.(2))) then raise Exit; Coqlib.check_required_library ["Coq";"Logic";"Eqdep_dec"]; - let new_eq_args = [|pf_unsafe_type_of gl ar1.(3);ar1.(3);ar2.(3)|] in + let new_eq_args = [|pf_unsafe_type_of gl (EConstr.of_constr ar1.(3));ar1.(3);ar2.(3)|] in let inj2 = Coqlib.coq_constant "inj_pair2_eq_dec is missing" ["Logic";"Eqdep_dec"] "inj_pair2_eq_dec" in let c, eff = find_scheme (!eq_dec_scheme_kind_name()) (Univ.out_punivs ind) in @@ -1339,7 +1340,7 @@ let inject_if_homogenous_dependent_pair ty = tclTHENS (cut (mkApp (ceq,new_eq_args))) [clear [destVar hyp]; Proofview.V82.tactic (Tacmach.refine - (mkApp(inj2,[|ar1.(0);mkConst c;ar1.(1);ar1.(2);ar1.(3);ar2.(3);hyp|]))) + (EConstr.of_constr (mkApp(inj2,[|ar1.(0);mkConst c;ar1.(1);ar1.(2);ar1.(3);ar2.(3);hyp|])))) ])] with Exit -> Proofview.tclUNIT () @@ -1384,7 +1385,7 @@ let inject_at_positions env sigma l2r (eq,_,(t,t1,t2)) eq_clause posns tac = (Proofview.tclIGNORE (Proofview.Monad.List.map (fun (pf,ty) -> tclTHENS (cut ty) [inject_if_homogenous_dependent_pair ty; - Proofview.V82.tactic (Tacmach.refine pf)]) + Proofview.V82.tactic (Tacmach.refine (EConstr.of_constr pf))]) (if l2r then List.rev injectors else injectors))) (tac (List.length injectors))) diff --git a/tactics/hipattern.ml b/tactics/hipattern.ml index 87e252a38..5d78fd585 100644 --- a/tactics/hipattern.ml +++ b/tactics/hipattern.ml @@ -437,7 +437,7 @@ let find_eq_data eqn = (* fails with PatternMatchingFailure *) let extract_eq_args gl = function | MonomorphicLeibnizEq (e1,e2) -> - let t = pf_unsafe_type_of gl e1 in (t,e1,e2) + let t = pf_unsafe_type_of gl (EConstr.of_constr e1) in (t,e1,e2) | PolymorphicLeibnizEq (t,e1,e2) -> (t,e1,e2) | HeterogenousEq (t1,e1,t2,e2) -> if pf_conv_x gl (EConstr.of_constr t1) (EConstr.of_constr t2) then (t1,e1,e2) @@ -463,6 +463,7 @@ let match_eq_nf gls eqn (ref, hetero) = let n = if hetero then 4 else 3 in let args = List.init n (fun i -> mkGPatVar ("X" ^ string_of_int (i + 1))) in let pat = mkPattern (mkGAppRef ref args) in + let eqn = EConstr.of_constr eqn in match Id.Map.bindings (pf_matches gls pat eqn) with | [(m1,t);(m2,x);(m3,y)] -> assert (Id.equal m1 meta1 && Id.equal m2 meta2 && Id.equal m3 meta3); diff --git a/tactics/inv.ml b/tactics/inv.ml index eebc67222..2f5186f81 100644 --- a/tactics/inv.ml +++ b/tactics/inv.ml @@ -438,7 +438,7 @@ let raw_inversion inv_kind id status names = let concl = Proofview.Goal.concl gl in let c = mkVar id in let (ind, t) = - try pf_apply Tacred.reduce_to_atomic_ind gl (pf_unsafe_type_of gl c) + try pf_apply Tacred.reduce_to_atomic_ind gl (EConstr.of_constr (pf_unsafe_type_of gl (EConstr.of_constr c))) with UserError _ -> let msg = str "The type of " ++ pr_id id ++ str " is not inductive." in CErrors.user_err msg diff --git a/tactics/tacticals.ml b/tactics/tacticals.ml index 2754db010..02909243d 100644 --- a/tactics/tacticals.ml +++ b/tactics/tacticals.ml @@ -628,7 +628,7 @@ module New = struct (Proofview.Goal.nf_enter { enter = begin fun gl -> let indclause = Tacmach.New.of_old (fun gl -> mk_clenv_from gl (c, t)) gl in (* applying elimination_scheme just a little modified *) - let elimclause = Tacmach.New.of_old (fun gls -> mk_clenv_from gls (elim,Tacmach.New.pf_unsafe_type_of gl elim)) gl in + let elimclause = Tacmach.New.of_old (fun gls -> mk_clenv_from gls (elim,Tacmach.New.pf_unsafe_type_of gl (EConstr.of_constr elim))) gl in let indmv = match kind_of_term (last_arg elimclause.evd (EConstr.of_constr elimclause.templval.Evd.rebus)) with | Meta mv -> mv @@ -678,7 +678,7 @@ module New = struct let elimination_then tac c = Proofview.Goal.nf_enter { enter = begin fun gl -> - let (ind,t) = pf_reduce_to_quantified_ind gl (pf_unsafe_type_of gl c) in + let (ind,t) = pf_reduce_to_quantified_ind gl (EConstr.of_constr (pf_unsafe_type_of gl (EConstr.of_constr c))) in let isrec,mkelim = match (Global.lookup_mind (fst (fst ind))).mind_record with | None -> true,gl_make_elim diff --git a/tactics/tactics.ml b/tactics/tactics.ml index 639a12b34..b9a219a2c 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -457,6 +457,7 @@ let assert_before_then_gen b naming t tac = let open Context.Rel.Declaration in Proofview.Goal.enter { enter = begin fun gl -> let id = find_name b (LocalAssum (Anonymous,t)) naming gl in + let t = EConstr.of_constr t in Tacticals.New.tclTHENLAST (Proofview.V82.tactic (fun gl -> @@ -476,6 +477,7 @@ let assert_after_then_gen b naming t tac = let open Context.Rel.Declaration in Proofview.Goal.enter { enter = begin fun gl -> let id = find_name b (LocalAssum (Anonymous,t)) naming gl in + let t = EConstr.of_constr t in Tacticals.New.tclTHENFIRST (Proofview.V82.tactic (fun gl -> @@ -1303,6 +1305,7 @@ let clenv_refine_in ?(sidecond_first=false) with_evars ?(with_classes=true) if not with_evars && occur_meta clenv.evd (EConstr.of_constr new_hyp_typ) then error_uninstantiated_metas new_hyp_typ clenv; let new_hyp_prf = clenv_value clenv in + let new_hyp_prf = EConstr.of_constr new_hyp_prf in let exact_tac = Proofview.V82.tactic (Tacmach.refine_no_check new_hyp_prf) in let naming = NamingMustBe (dloc,targetid) in let with_clear = do_replace (Some id) naming in @@ -1434,7 +1437,7 @@ let general_elim with_evars clear_flag (c, lbindc) elim = let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in let ct = Retyping.get_type_of env sigma (EConstr.of_constr c) in - let t = try snd (reduce_to_quantified_ind env sigma ct) with UserError _ -> ct in + let t = try snd (reduce_to_quantified_ind env sigma (EConstr.of_constr ct)) with UserError _ -> ct in let elimtac = elimination_clause_scheme with_evars in let indclause = make_clenv_binding env sigma (c, t) lbindc in let sigma = meta_merge sigma (clear_metas indclause.evd) in @@ -1452,6 +1455,7 @@ let general_case_analysis_in_context with_evars clear_flag (c,lbindc) = let env = Proofview.Goal.env gl in let concl = Proofview.Goal.concl gl in let t = Retyping.get_type_of env (Sigma.to_evar_map sigma) (EConstr.of_constr c) in + let t = EConstr.of_constr t in let (mind,_) = reduce_to_quantified_ind env (Sigma.to_evar_map sigma) t in let sort = Tacticals.New.elimination_sort_of_goal gl in let Sigma (elim, sigma, p) = @@ -1491,7 +1495,8 @@ let find_ind_eliminator ind s gl = evd, c let find_eliminator c gl = - let ((ind,u),t) = Tacmach.New.pf_reduce_to_quantified_ind gl (Tacmach.New.pf_unsafe_type_of gl c) in + let c = EConstr.of_constr c in + let ((ind,u),t) = Tacmach.New.pf_reduce_to_quantified_ind gl (EConstr.of_constr (Tacmach.New.pf_unsafe_type_of gl c)) in if is_nonrec ind then raise IsNonrec; let evd, c = find_ind_eliminator ind (Tacticals.New.elimination_sort_of_goal gl) gl in evd, {elimindex = None; elimbody = (c,NoBindings); @@ -1637,6 +1642,7 @@ let descend_in_conjunctions avoid tac (err, info) c = let sigma = Tacmach.New.project gl in try let t = Retyping.get_type_of env sigma (EConstr.of_constr c) in + let t = EConstr.of_constr t in let ((ind,u),t) = reduce_to_quantified_ind env sigma t in let sign,ccl = decompose_prod_assum t in match match_with_tuple sigma ccl with @@ -1661,6 +1667,7 @@ let descend_in_conjunctions avoid tac (err, info) c = match make_projection env sigma params cstr sign elim i n c u with | None -> Tacticals.New.tclFAIL 0 (mt()) | Some (p,pt) -> + let p = EConstr.of_constr p in Tacticals.New.tclTHENS (assert_before_gen false (NamingAvoid avoid) pt) [Proofview.V82.tactic (refine p); @@ -1920,7 +1927,7 @@ let apply_in_delayed_once sidecond_first with_delta with_destruct with_evars nam let cut_and_apply c = Proofview.Goal.nf_enter { enter = begin fun gl -> let sigma = Tacmach.New.project gl in - match kind_of_term (Tacmach.New.pf_hnf_constr gl (Tacmach.New.pf_unsafe_type_of gl c)) with + match kind_of_term (Tacmach.New.pf_hnf_constr gl (EConstr.of_constr (Tacmach.New.pf_unsafe_type_of gl (EConstr.of_constr c)))) with | Prod (_,c1,c2) when EConstr.Vars.noccurn sigma 1 (EConstr.of_constr c2) -> let concl = Proofview.Goal.concl gl in let env = Tacmach.New.pf_env gl in @@ -2201,6 +2208,7 @@ let constructor_tac with_evars expctdnumopt i lbind = let reduce_to_quantified_ind = Tacmach.New.pf_apply Tacred.reduce_to_quantified_ind gl in + let cl = EConstr.of_constr cl in let (mind,redcl) = reduce_to_quantified_ind cl in let nconstr = Array.length (snd (Global.lookup_inductive (fst mind))).mind_consnames in @@ -2240,6 +2248,7 @@ let any_constructor with_evars tacopt = let reduce_to_quantified_ind = Tacmach.New.pf_apply Tacred.reduce_to_quantified_ind gl in + let cl = EConstr.of_constr cl in let mind = fst (reduce_to_quantified_ind cl) in let nconstr = Array.length (snd (Global.lookup_inductive (fst mind))).mind_consnames in @@ -2298,7 +2307,8 @@ let my_find_eq_data_decompose gl t = let intro_decomp_eq loc l thin tac id = Proofview.Goal.nf_enter { enter = begin fun gl -> let c = mkVar id in - let t = Tacmach.New.pf_unsafe_type_of gl c in + let t = Tacmach.New.pf_unsafe_type_of gl (EConstr.of_constr c) in + let t = EConstr.of_constr t in let _,t = Tacmach.New.pf_reduce_to_quantified_ind gl t in match my_find_eq_data_decompose gl t with | Some (eq,u,eq_args) -> @@ -2312,7 +2322,8 @@ let intro_decomp_eq loc l thin tac id = let intro_or_and_pattern loc with_evars bracketed ll thin tac id = Proofview.Goal.enter { enter = begin fun gl -> let c = mkVar id in - let t = Tacmach.New.pf_unsafe_type_of gl c in + let t = Tacmach.New.pf_unsafe_type_of gl (EConstr.of_constr c) in + let t = EConstr.of_constr t in let (ind,t) = Tacmach.New.pf_reduce_to_quantified_ind gl t in let branchsigns = compute_constructor_signatures false ind in let nv_with_let = Array.map List.length branchsigns in @@ -2337,7 +2348,7 @@ let rewrite_hyp_then assert_style with_evars thin l2r id tac = let sigma = Tacmach.New.project gl in let type_of = Tacmach.New.pf_unsafe_type_of gl in let whd_all = Tacmach.New.pf_apply whd_all gl in - let t = whd_all (EConstr.of_constr (type_of (mkVar id))) in + let t = whd_all (EConstr.of_constr (type_of (EConstr.mkVar id))) in let eqtac, thin = match match_with_equality_type sigma t with | Some (hdcncl,[_;lhs;rhs]) -> if l2r && isVar lhs && not (occur_var env sigma (destVar lhs) (EConstr.of_constr rhs)) then @@ -2747,7 +2758,7 @@ let forward b usetac ipat c = match usetac with | None -> Proofview.Goal.enter { enter = begin fun gl -> - let t = Tacmach.New.pf_unsafe_type_of gl c in + let t = Tacmach.New.pf_unsafe_type_of gl (EConstr.of_constr c) in let hd = head_ident c in Tacticals.New.tclTHENFIRST (assert_as true hd ipat t) (exact_no_check c) end } @@ -3233,7 +3244,7 @@ let atomize_param_of_ind_then (indref,nparams,_) hyp0 tac = | Var id -> id | _ -> let type_of = Tacmach.New.pf_unsafe_type_of gl in - id_of_name_using_hdchar (Global.env()) (type_of c) Anonymous in + id_of_name_using_hdchar (Global.env()) (type_of (EConstr.of_constr c)) Anonymous in let x = fresh_id_in_env avoid id env in Tacticals.New.tclTHEN (letin_tac None (Name x) c None allHypsAndConcl) @@ -3645,7 +3656,7 @@ let abstract_args gl generalize_vars dep id defined f args = let decl = List.hd rel in RelDecl.get_name decl, RelDecl.get_type decl, c in - let argty = Tacmach.pf_unsafe_type_of gl arg in + let argty = Tacmach.pf_unsafe_type_of gl (EConstr.of_constr arg) in let sigma', ty = Evarsolve.refresh_universes (Some true) env !sigma (EConstr.of_constr ty) in let () = sigma := sigma' in let lenctx = List.length ctx in @@ -3686,7 +3697,7 @@ let abstract_args gl generalize_vars dep id defined f args = true, mkApp (f', before), after in if dogen then - let tyf' = Tacmach.pf_unsafe_type_of gl f' in + let tyf' = Tacmach.pf_unsafe_type_of gl (EConstr.of_constr f') in let arity, ctx, ctxenv, c', args, eqs, refls, nogen, vars, env = Array.fold_left aux (tyf',[],env,f',[],[],[],Id.Set.empty,Id.Set.empty,env) args' in @@ -3794,6 +3805,7 @@ let specialize_eqs id gl = let ty' = Tacred.whd_simpl env !evars (EConstr.of_constr ty') and acc' = Tacred.whd_simpl env !evars (EConstr.of_constr acc') in let ty' = Evarutil.nf_evar !evars ty' in + let ty' = EConstr.of_constr ty' in if worked then tclTHENFIRST (Tacmach.internal_cut true id ty') (Proofview.V82.of_tactic (exact_no_check ((* refresh_universes_strict *) acc'))) gl @@ -4014,6 +4026,7 @@ let compute_elim_signature (evd,(elimc,elimt),ind_type_guess) names_info = let guess_elim isrec dep s hyp0 gl = let tmptyp0 = Tacmach.New.pf_get_hyp_typ hyp0 gl in + let tmptyp0 = EConstr.of_constr tmptyp0 in let mind,_ = Tacmach.New.pf_reduce_to_quantified_ind gl tmptyp0 in let evd, elimc = if isrec && not (is_nonrec (fst mind)) then find_ind_eliminator (fst mind) s gl @@ -4028,12 +4041,13 @@ let guess_elim isrec dep s hyp0 gl = let Sigma (ind, sigma, _) = build_case_analysis_scheme_default env sigma mind s in (Sigma.to_evar_map sigma, ind) in - let elimt = Tacmach.New.pf_unsafe_type_of gl elimc in + let elimt = Tacmach.New.pf_unsafe_type_of gl (EConstr.of_constr elimc) in evd, ((elimc, NoBindings), elimt), mkIndU mind let given_elim hyp0 (elimc,lbind as e) gl = let tmptyp0 = Tacmach.New.pf_get_hyp_typ hyp0 gl in let ind_type_guess,_ = decompose_app ((strip_prod tmptyp0)) in + let elimc = EConstr.of_constr elimc in Tacmach.New.project gl, (e, Tacmach.New.pf_unsafe_type_of gl elimc), ind_type_guess type scheme_signature = @@ -4069,7 +4083,7 @@ let get_elim_signature elim hyp0 gl = compute_elim_signature (given_elim hyp0 elim gl) hyp0 let is_functional_induction elimc gl = - let scheme = compute_elim_sig ~elimc (Tacmach.New.pf_unsafe_type_of gl (fst elimc)) in + let scheme = compute_elim_sig ~elimc (Tacmach.New.pf_unsafe_type_of gl (EConstr.of_constr (fst elimc))) in (* The test is not safe: with non-functional induction on non-standard induction scheme, this may fail *) Option.is_empty scheme.indarg @@ -4466,7 +4480,7 @@ let induction_gen_l isrec with_evars elim names lc = let type_of = Tacmach.New.pf_unsafe_type_of gl in let sigma = Tacmach.New.project gl in let x = - id_of_name_using_hdchar (Global.env()) (type_of c) Anonymous in + id_of_name_using_hdchar (Global.env()) (type_of (EConstr.of_constr c)) Anonymous in let id = new_fresh_id [] x gl in let newl' = List.map (fun r -> replace_term sigma (EConstr.of_constr c) (EConstr.mkVar id) (EConstr.of_constr r)) l' in @@ -4606,6 +4620,7 @@ let elim_scheme_type elim t = let elim_type t = Proofview.Goal.s_enter { s_enter = begin fun gl -> + let t = EConstr.of_constr t in let (ind,t) = Tacmach.New.pf_apply reduce_to_atomic_ind gl t in let evd, elimc = find_ind_eliminator (fst ind) (Tacticals.New.elimination_sort_of_goal gl) gl in Sigma.Unsafe.of_pair (elim_scheme_type elimc t, evd) @@ -4613,6 +4628,7 @@ let elim_type t = let case_type t = Proofview.Goal.s_enter { s_enter = begin fun gl -> + let t = EConstr.of_constr t in let sigma = Proofview.Goal.sigma gl in let env = Tacmach.New.pf_env gl in let (ind,t) = reduce_to_atomic_ind env (Sigma.to_evar_map sigma) t in @@ -4717,7 +4733,7 @@ let (forward_setoid_symmetry_in, setoid_symmetry_in) = Hook.make () let symmetry_in id = Proofview.Goal.enter { enter = begin fun gl -> - let ctype = Tacmach.New.pf_unsafe_type_of gl (mkVar id) in + let ctype = Tacmach.New.pf_unsafe_type_of gl (EConstr.mkVar id) in let sign,t = decompose_prod_assum ctype in Proofview.tclORELSE begin diff --git a/toplevel/auto_ind_decl.ml b/toplevel/auto_ind_decl.ml index 02c43aec5..6561627f6 100644 --- a/toplevel/auto_ind_decl.ml +++ b/toplevel/auto_ind_decl.ml @@ -363,7 +363,7 @@ let do_replace_lb mode lb_scheme_key aavoid narg p q = ) in Proofview.Goal.nf_enter { enter = begin fun gl -> - let type_of_pq = Tacmach.New.of_old (fun gl -> pf_unsafe_type_of gl p) gl in + let type_of_pq = Tacmach.New.of_old (fun gl -> pf_unsafe_type_of gl (EConstr.of_constr p)) gl in let u,v = destruct_ind type_of_pq in let lb_type_of_p = try @@ -425,7 +425,7 @@ let do_replace_bl mode bl_scheme_key (ind,u as indu) aavoid narg lft rgt = match (l1,l2) with | (t1::q1,t2::q2) -> Proofview.Goal.enter { enter = begin fun gl -> - let tt1 = Tacmach.New.pf_unsafe_type_of gl t1 in + let tt1 = Tacmach.New.pf_unsafe_type_of gl (EConstr.of_constr t1) in if eq_constr t1 t2 then aux q1 q2 else ( let u,v = try destruct_ind tt1 diff --git a/toplevel/command.ml b/toplevel/command.ml index dbf256ba8..80f3b26e4 100644 --- a/toplevel/command.ml +++ b/toplevel/command.ml @@ -1164,7 +1164,7 @@ let declare_fixpoint local poly ((fixnames,fixdefs,fixtypes),pl,ctx,fiximps) ind List.map3 (fun id t (len,imps,_) -> ((id,pl),(t,(len,imps)))) fixnames fixtypes fiximps in let init_tac = - Some (List.map (Option.cata Tacmach.refine_no_check Tacticals.tclIDTAC) + Some (List.map (Option.cata (EConstr.of_constr %> Tacmach.refine_no_check) Tacticals.tclIDTAC) fixdefs) in let init_tac = Option.map (List.map Proofview.V82.tactic) init_tac @@ -1201,7 +1201,7 @@ let declare_cofixpoint local poly ((fixnames,fixdefs,fixtypes),pl,ctx,fiximps) n List.map3 (fun id t (len,imps,_) -> ((id,pl),(t,(len,imps)))) fixnames fixtypes fiximps in let init_tac = - Some (List.map (Option.cata Tacmach.refine_no_check Tacticals.tclIDTAC) + Some (List.map (Option.cata (EConstr.of_constr %> Tacmach.refine_no_check) Tacticals.tclIDTAC) fixdefs) in let init_tac = Option.map (List.map Proofview.V82.tactic) init_tac -- cgit v1.2.3 From 0489e8b56d7e10f7111c0171960e25d32201b963 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 11 Nov 2016 21:55:33 +0100 Subject: Clenv API using EConstr. --- engine/eConstr.ml | 2 + engine/eConstr.mli | 3 ++ ltac/rewrite.ml | 10 ++-- pretyping/evarconv.ml | 12 ++--- pretyping/miscops.ml | 13 +++++ pretyping/miscops.mli | 5 ++ pretyping/reductionops.ml | 8 +-- pretyping/reductionops.mli | 4 +- pretyping/typing.ml | 3 +- pretyping/typing.mli | 2 +- pretyping/unification.ml | 14 +++-- proofs/clenv.ml | 129 ++++++++++++++++++++++++--------------------- proofs/clenv.mli | 31 ++++++----- proofs/clenvtac.ml | 17 +++--- proofs/clenvtac.mli | 1 + tactics/auto.ml | 5 +- tactics/autorewrite.ml | 6 +-- tactics/class_tactics.ml | 21 +++++--- tactics/eauto.ml | 3 +- tactics/equality.ml | 38 ++++++++----- tactics/hints.ml | 12 +++-- tactics/leminv.ml | 4 +- tactics/tacticals.ml | 12 +++-- tactics/tactics.ml | 70 ++++++++++++++++-------- 24 files changed, 261 insertions(+), 164 deletions(-) diff --git a/engine/eConstr.ml b/engine/eConstr.ml index 095521e25..7bd708e31 100644 --- a/engine/eConstr.ml +++ b/engine/eConstr.ml @@ -560,6 +560,8 @@ 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 diff --git a/engine/eConstr.mli b/engine/eConstr.mli index 10eb064a3..e4136a612 100644 --- a/engine/eConstr.mli +++ b/engine/eConstr.mli @@ -171,6 +171,9 @@ 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 + end (** {5 Unsafe operations} *) diff --git a/ltac/rewrite.ml b/ltac/rewrite.ml index 37b9a9edb..52cf1ff35 100644 --- a/ltac/rewrite.ml +++ b/ltac/rewrite.ml @@ -507,9 +507,12 @@ let decompose_applied_relation env sigma (c,l) = let open Context.Rel.Declaration in let ctype = Retyping.get_type_of env sigma (EConstr.of_constr c) in let find_rel ty = + let ty = EConstr.of_constr ty in let sigma, cl = Clenv.make_evar_clause env sigma ty in + let l = Miscops.map_bindings EConstr.of_constr l in let sigma = Clenv.solve_evar_clause env sigma true cl l in let { Clenv.cl_holes = holes; Clenv.cl_concl = t } = cl in + let t = EConstr.Unsafe.to_constr t in let (equiv, c1, c2) = decompose_app_rel env sigma t in let ty1 = Retyping.get_type_of env sigma (EConstr.of_constr c1) in let ty2 = Retyping.get_type_of env sigma (EConstr.of_constr c2) in @@ -517,7 +520,7 @@ let decompose_applied_relation env sigma (c,l) = | None -> None | Some sigma -> let sort = sort_of_rel env sigma equiv in - let args = Array.map_of_list (fun h -> h.Clenv.hole_evar) holes in + let args = Array.map_of_list (fun h -> EConstr.Unsafe.to_constr h.Clenv.hole_evar) holes in let value = mkApp (c, args) in Some (sigma, { prf=value; car=ty1; rel = equiv; sort = Sorts.is_prop sort; @@ -618,9 +621,10 @@ let solve_remaining_by env sigma holes by = | Some tac -> let map h = if h.Clenv.hole_deps then None - else - let (evk, _) = destEvar (h.Clenv.hole_evar) in + else match EConstr.kind sigma h.Clenv.hole_evar with + | Evar (evk, _) -> Some evk + | _ -> None in (** Only solve independent holes *) let indep = List.map_filter map holes in diff --git a/pretyping/evarconv.ml b/pretyping/evarconv.ml index 77e91095f..ee6355736 100644 --- a/pretyping/evarconv.ml +++ b/pretyping/evarconv.ml @@ -6,8 +6,6 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -module CVars = Vars - open CErrors open Util open Names @@ -184,10 +182,12 @@ let check_conv_record env sigma (t1,sk1) (t2,sk2) = | None -> raise Not_found | Some (l',el,s') -> (l'@Stack.append_app [|el|] Stack.empty,s') in let subst, ctx' = Universes.fresh_universe_context_set_instance ctx in - let c' = EConstr.of_constr (CVars.subst_univs_level_constr subst c) in - let t' = CVars.subst_univs_level_constr subst t' in - let bs' = List.map (CVars.subst_univs_level_constr subst %> EConstr.of_constr) bs in - let h, _ = decompose_app_vect sigma (EConstr.of_constr t') in + let c = EConstr.of_constr c in + let c' = subst_univs_level_constr subst c in + let t' = EConstr.of_constr t' in + let t' = subst_univs_level_constr subst t' in + let bs' = List.map (EConstr.of_constr %> subst_univs_level_constr subst) bs in + let h, _ = decompose_app_vect sigma t' in ctx',(EConstr.of_constr h, t2),c',bs',(Stack.append_app_list params Stack.empty,params1), (Stack.append_app_list us Stack.empty,us2),(extra_args1,extra_args2),c1, (n, Stack.zip sigma (t2,sk2)) diff --git a/pretyping/miscops.ml b/pretyping/miscops.ml index 142e430ff..7fe81c9a4 100644 --- a/pretyping/miscops.ml +++ b/pretyping/miscops.ml @@ -58,3 +58,16 @@ let map_red_expr_gen f g h = function | CbvNative occs_o -> CbvNative (Option.map (map_occs (map_union g h)) occs_o) | Cbn flags -> Cbn (map_flags g flags) | ExtraRedExpr _ | Red _ | Hnf as x -> x + +(** Mapping bindings *) + +let map_explicit_bindings f l = + let map (loc, hyp, x) = (loc, hyp, f x) in + List.map map l + +let map_bindings f = function +| ImplicitBindings l -> ImplicitBindings (List.map f l) +| ExplicitBindings expl -> ExplicitBindings (map_explicit_bindings f expl) +| NoBindings -> NoBindings + +let map_with_bindings f (x, bl) = (f x, map_bindings f bl) diff --git a/pretyping/miscops.mli b/pretyping/miscops.mli index 337473a6f..f30dc1a4b 100644 --- a/pretyping/miscops.mli +++ b/pretyping/miscops.mli @@ -27,3 +27,8 @@ val intro_pattern_naming_eq : val map_red_expr_gen : ('a -> 'd) -> ('b -> 'e) -> ('c -> 'f) -> ('a,'b,'c) red_expr_gen -> ('d,'e,'f) red_expr_gen + +(** Mapping bindings *) + +val map_bindings : ('a -> 'b) -> 'a bindings -> 'b bindings +val map_with_bindings : ('a -> 'b) -> 'a with_bindings -> 'b with_bindings diff --git a/pretyping/reductionops.ml b/pretyping/reductionops.ml index 0e0b80744..0dd615bfb 100644 --- a/pretyping/reductionops.ml +++ b/pretyping/reductionops.ml @@ -1588,9 +1588,11 @@ let meta_instance sigma b = if Metaset.is_empty fm then b.rebus else let c_sigma = Metamap.bind (fun mv -> meta_value sigma mv) fm in - instance sigma c_sigma (EConstr.of_constr b.rebus) + EConstr.of_constr (instance sigma c_sigma b.rebus) -let nf_meta sigma c = meta_instance sigma (mk_freelisted c) +let nf_meta sigma c = + let cl = mk_freelisted c in + EConstr.Unsafe.to_constr (meta_instance sigma { cl with rebus = EConstr.of_constr cl.rebus }) (* Instantiate metas that create beta/iota redexes *) @@ -1652,7 +1654,7 @@ let meta_reducible_instance evd b = | _ -> EConstr.map evd irec u in if Metaset.is_empty fm then (* nf_betaiota? *) b.rebus - else EConstr.Unsafe.to_constr (irec (EConstr.of_constr b.rebus)) + else irec b.rebus let head_unfold_under_prod ts env sigma c = diff --git a/pretyping/reductionops.mli b/pretyping/reductionops.mli index c3b82729d..864b1a625 100644 --- a/pretyping/reductionops.mli +++ b/pretyping/reductionops.mli @@ -295,6 +295,6 @@ val whd_betaiota_deltazeta_for_iota_state : state * Cst_stack.t (** {6 Meta-related reduction functions } *) -val meta_instance : evar_map -> constr freelisted -> constr +val meta_instance : evar_map -> EConstr.constr freelisted -> EConstr.constr val nf_meta : evar_map -> constr -> constr -val meta_reducible_instance : evar_map -> constr freelisted -> constr +val meta_reducible_instance : evar_map -> EConstr.constr freelisted -> EConstr.constr diff --git a/pretyping/typing.ml b/pretyping/typing.ml index cf58a0b66..29697260f 100644 --- a/pretyping/typing.ml +++ b/pretyping/typing.ml @@ -40,6 +40,7 @@ let meta_type evd mv = let ty = try Evd.meta_ftype evd mv with Not_found -> anomaly (str "unknown meta ?" ++ str (Nameops.string_of_meta mv)) in + let ty = Evd.map_fl EConstr.of_constr ty in meta_instance evd ty let constant_type_knowing_parameters env sigma cst jl = @@ -256,7 +257,7 @@ let rec execute env evdref cstr = let cstr = EConstr.of_constr (whd_evar !evdref (EConstr.Unsafe.to_constr cstr)) in match EConstr.kind !evdref cstr with | Meta n -> - { uj_val = cstr; uj_type = EConstr.of_constr (meta_type !evdref n) } + { uj_val = cstr; uj_type = meta_type !evdref n } | Evar ev -> let ty = EConstr.existential_type !evdref ev in diff --git a/pretyping/typing.mli b/pretyping/typing.mli index 1fb414906..94a56b6e1 100644 --- a/pretyping/typing.mli +++ b/pretyping/typing.mli @@ -31,7 +31,7 @@ val e_sort_of : env -> evar_map ref -> EConstr.types -> sorts val e_check : env -> evar_map ref -> EConstr.constr -> EConstr.types -> unit (** Returns the instantiated type of a metavariable *) -val meta_type : evar_map -> metavariable -> types +val meta_type : evar_map -> metavariable -> EConstr.types (** Solve existential variables using typing *) val e_solve_evars : env -> evar_map ref -> EConstr.constr -> constr diff --git a/pretyping/unification.ml b/pretyping/unification.ml index 483fefaf2..2b2069ec4 100644 --- a/pretyping/unification.ml +++ b/pretyping/unification.ml @@ -617,10 +617,10 @@ let subst_defined_metas_evars sigma (bl,el) c = in try Some (substrec c) with Not_found -> None let check_compatibility env pbty flags (sigma,metasubst,evarsubst : subst0) tyM tyN = - match subst_defined_metas_evars sigma (metasubst,[]) (EConstr.of_constr tyM) with + match subst_defined_metas_evars sigma (metasubst,[]) tyM with | None -> sigma | Some m -> - match subst_defined_metas_evars sigma (metasubst,[]) (EConstr.of_constr tyN) with + match subst_defined_metas_evars sigma (metasubst,[]) tyN with | None -> sigma | Some n -> if is_ground_term sigma m && is_ground_term sigma n then @@ -705,6 +705,7 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst : subst0) conv_at_top e (try let tyM = Typing.meta_type sigma k in let tyN = get_type_of curenv ~lax:true sigma cN in + let tyN = EConstr.of_constr tyN in check_compatibility curenv CUMUL flags substn tyN tyM with RetypeError _ -> (* Renounce, maybe metas/evars prevents typing *) sigma) @@ -724,6 +725,7 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst : subst0) conv_at_top e if opt.with_types && flags.check_applied_meta_types then (try let tyM = get_type_of curenv ~lax:true sigma cM in + let tyM = EConstr.of_constr tyM in let tyN = Typing.meta_type sigma k in check_compatibility curenv CUMUL flags substn tyM tyN with RetypeError _ -> @@ -977,6 +979,8 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst : subst0) conv_at_top e try (* Ensure we call conversion on terms of the same type *) let tyM = get_type_of curenv ~lax:true sigma m1 in let tyN = get_type_of curenv ~lax:true sigma n1 in + let tyM = EConstr.of_constr tyM in + let tyN = EConstr.of_constr tyN in check_compatibility curenv CUMUL flags substn tyM tyN with RetypeError _ -> (* Renounce, maybe metas/evars prevents typing *) sigma @@ -1265,7 +1269,7 @@ let w_coerce_to_type env evd c cty mvty = let w_coerce env evd mv c = let cty = get_type_of env evd c in let mvty = Typing.meta_type evd mv in - w_coerce_to_type env evd c (EConstr.of_constr cty) (EConstr.of_constr mvty) + w_coerce_to_type env evd c (EConstr.of_constr cty) mvty let unify_to_type env sigma flags c status u = let sigma, c = refresh_universes (Some false) env sigma c in @@ -1275,6 +1279,7 @@ let unify_to_type env sigma flags c status u = let unify_type env sigma flags mv status c = let mvty = Typing.meta_type sigma mv in + let mvty = EConstr.Unsafe.to_constr mvty in let mvty = nf_meta sigma mvty in unify_to_type env sigma (set_flags_for_type flags) @@ -1923,7 +1928,6 @@ let secondOrderAbstraction env evd flags typ (p, oplist) = let flags = { flags with core_unify_flags = flags.subterm_unify_flags } in let (evd',cllist) = w_unify_to_subterm_list env evd flags p oplist typ in let typp = Typing.meta_type evd' p in - let typp = EConstr.of_constr typp in let evd',(pred,predtyp) = abstract_list_all env evd' typp typ cllist in let evd', b = infer_conv ~pb:CUMUL env evd' predtyp typp in if not b then @@ -1942,7 +1946,7 @@ let secondOrderAbstraction env evd flags typ (p, oplist) = let secondOrderDependentAbstraction env evd flags typ (p, oplist) = let typp = Typing.meta_type evd p in - let evd, pred = abstract_list_all_with_dependencies env evd (EConstr.of_constr typp) typ oplist in + let evd, pred = abstract_list_all_with_dependencies env evd typp typ oplist in let pred = EConstr.of_constr pred in w_merge env false flags.merge_unify_flags (evd,[p,pred,(Conv,TypeProcessed)],[]) diff --git a/proofs/clenv.ml b/proofs/clenv.ml index 67ed5daaa..3e3889cbb 100644 --- a/proofs/clenv.ml +++ b/proofs/clenv.ml @@ -12,11 +12,12 @@ open Util open Names open Nameops open Term -open Vars open Termops open Namegen open Environ open Evd +open EConstr +open Vars open Reduction open Reductionops open Tacred @@ -29,7 +30,7 @@ open Sigma.Notations (* Abbreviations *) let pf_env = Refiner.pf_env -let pf_type_of gls c = Typing.unsafe_type_of (pf_env gls) gls.sigma (EConstr.of_constr c) +let pf_type_of gls c = Typing.unsafe_type_of (pf_env gls) gls.sigma c (******************************************************************) (* Clausal environments *) @@ -43,12 +44,6 @@ type clausenv = { let cl_env ce = ce.env let cl_sigma ce = ce.evd -let map_clenv sub clenv = - { templval = map_fl sub clenv.templval; - templtyp = map_fl sub clenv.templtyp; - evd = cmap sub clenv.evd; - env = clenv.env } - let clenv_nf_meta clenv c = nf_meta clenv.evd c let clenv_term clenv c = meta_instance clenv.evd c let clenv_meta_type clenv mv = Typing.meta_type clenv.evd mv @@ -56,30 +51,34 @@ let clenv_value clenv = meta_instance clenv.evd clenv.templval let clenv_type clenv = meta_instance clenv.evd clenv.templtyp let refresh_undefined_univs clenv = - match kind_of_term clenv.templval.rebus with + match EConstr.kind clenv.evd clenv.templval.rebus with | Var _ -> clenv, Univ.empty_level_subst - | App (f, args) when isVar f -> clenv, Univ.empty_level_subst + | App (f, args) when isVar clenv.evd f -> clenv, Univ.empty_level_subst | _ -> let evd', subst = Evd.refresh_undefined_universes clenv.evd in let map_freelisted f = { f with rebus = subst_univs_level_constr subst f.rebus } in { clenv with evd = evd'; templval = map_freelisted clenv.templval; templtyp = map_freelisted clenv.templtyp }, subst -let clenv_hnf_constr ce t = hnf_constr (cl_env ce) (cl_sigma ce) t +let clenv_hnf_constr ce t = EConstr.of_constr (hnf_constr (cl_env ce) (cl_sigma ce) t) -let clenv_get_type_of ce c = Retyping.get_type_of (cl_env ce) (cl_sigma ce) (EConstr.of_constr c) +let clenv_get_type_of ce c = EConstr.of_constr (Retyping.get_type_of (cl_env ce) (cl_sigma ce) c) exception NotExtensibleClause +let mk_freelisted c = + map_fl EConstr.of_constr (mk_freelisted (EConstr.Unsafe.to_constr c)) + let clenv_push_prod cl = - let typ = whd_all (cl_env cl) (cl_sigma cl) (EConstr.of_constr (clenv_type cl)) in - let rec clrec typ = match kind_of_term typ with + let typ = whd_all (cl_env cl) (cl_sigma cl) (clenv_type cl) in + let typ = EConstr.of_constr typ in + let rec clrec typ = match EConstr.kind cl.evd typ with | Cast (t,_,_) -> clrec t | Prod (na,t,u) -> let mv = new_meta () in - let dep = not (EConstr.Vars.noccurn (cl_sigma cl) 1 (EConstr.of_constr u)) in + let dep = not (noccurn (cl_sigma cl) 1 u) in let na' = if dep then na else Anonymous in - let e' = meta_declare mv t ~name:na' cl.evd in + let e' = meta_declare mv (EConstr.Unsafe.to_constr t) ~name:na' cl.evd in let concl = if dep then subst1 (mkMeta mv) u else u in let def = applist (cl.templval.rebus,[mkMeta mv]) in { templval = mk_freelisted def; @@ -103,14 +102,17 @@ let clenv_push_prod cl = and [ccl] is [forall y, Meta n1=y -> y=Meta n1] *) let clenv_environments evd bound t = + let open EConstr in + let open Vars in let rec clrec (e,metas) n t = - match n, kind_of_term t with + match n, EConstr.kind evd t with | (Some 0, _) -> (e, List.rev metas, t) | (n, Cast (t,_,_)) -> clrec (e,metas) n t | (n, Prod (na,t1,t2)) -> let mv = new_meta () in - let dep = not (noccurn 1 t2) in + let dep = not (noccurn evd 1 t2) in let na' = if dep then na else Anonymous in + let t1 = EConstr.Unsafe.to_constr t1 in let e' = meta_declare mv t1 ~name:na' e in clrec (e', (mkMeta mv)::metas) (Option.map ((+) (-1)) n) (if dep then (subst1 (mkMeta mv) t2) else t2) @@ -132,7 +134,7 @@ let mk_clenv_from_n gls n (c,cty) = let mk_clenv_from gls = mk_clenv_from_n gls None -let mk_clenv_type_of gls t = mk_clenv_from gls (t,pf_type_of gls t) +let mk_clenv_type_of gls t = mk_clenv_from gls (t,EConstr.of_constr (pf_type_of gls t)) (******************************************************************) @@ -168,13 +170,13 @@ let clenv_assign mv rhs clenv = error "clenv_assign: circularity in unification"; try if meta_defined clenv.evd mv then - if not (Term.eq_constr (fst (meta_fvalue clenv.evd mv)).rebus rhs) then + if not (EConstr.eq_constr clenv.evd (EConstr.of_constr (fst (meta_fvalue clenv.evd mv)).rebus) rhs) then error_incompatible_inst clenv mv else clenv else let st = (Conv,TypeNotProcessed) in - {clenv with evd = meta_assign mv (rhs_fls.rebus,st) clenv.evd} + {clenv with evd = meta_assign mv (EConstr.Unsafe.to_constr rhs_fls.rebus,st) clenv.evd} with Not_found -> error "clenv_assign: undefined meta" @@ -219,7 +221,7 @@ let clenv_assign mv rhs clenv = *) let clenv_metas_in_type_of_meta evd mv = - (mk_freelisted (meta_instance evd (meta_ftype evd mv))).freemetas + (mk_freelisted (meta_instance evd (map_fl EConstr.of_constr (meta_ftype evd mv)))).freemetas let dependent_in_type_of_metas clenv mvs = List.fold_right @@ -257,15 +259,14 @@ let clenv_dependent ce = clenv_dependent_gen false ce let clenv_unify ?(flags=default_unify_flags ()) cv_pb t1 t2 clenv = { clenv with - evd = w_unify ~flags clenv.env clenv.evd cv_pb (EConstr.of_constr t1) (EConstr.of_constr t2) } + evd = w_unify ~flags clenv.env clenv.evd cv_pb t1 t2 } let clenv_unify_meta_types ?(flags=default_unify_flags ()) clenv = { clenv with evd = w_unify_meta_types ~flags:flags clenv.env clenv.evd } let clenv_unique_resolver ?(flags=default_unify_flags ()) clenv gl = let concl = Goal.V82.concl clenv.evd (sig_it gl) in - let concl = EConstr.Unsafe.to_constr concl in - if isMeta (fst (decompose_appvect (whd_nored clenv.evd (EConstr.of_constr clenv.templtyp.rebus)))) then + if isMeta clenv.evd (EConstr.of_constr (fst (decompose_app_vect clenv.evd (EConstr.of_constr (whd_nored clenv.evd clenv.templtyp.rebus))))) then clenv_unify CUMUL ~flags (clenv_type clenv) concl (clenv_unify_meta_types ~flags clenv) else @@ -275,23 +276,19 @@ let clenv_unique_resolver ?(flags=default_unify_flags ()) clenv gl = let adjust_meta_source evd mv = function | loc,Evar_kinds.VarInstance id -> let rec match_name c l = - match kind_of_term c, l with - | Lambda (Name id,_,c), a::l when Constr.equal a (mkMeta mv) -> Some id + match EConstr.kind evd c, l with + | Lambda (Name id,_,c), a::l when EConstr.eq_constr evd a (mkMeta mv) -> Some id | Lambda (_,_,c), a::l -> match_name c l | _ -> None in (* This is very ad hoc code so that an evar inherits the name of the binder in situations like "ex_intro (fun x => P) ?ev p" *) let f = function (mv',(Cltyp (_,t) | Clval (_,_,t))) -> if Metaset.mem mv t.freemetas then - let f,l = decompose_app t.rebus in - match kind_of_term f with + let f,l = decompose_app evd (EConstr.of_constr t.rebus) in + match EConstr.kind evd f with | Meta mv'' -> (match meta_opt_fvalue evd mv'' with - | Some (c,_) -> match_name c.rebus l - | None -> None) - | Evar ev -> - (match existential_opt_value evd ev with - | Some c -> match_name c l + | Some (c,_) -> match_name (EConstr.of_constr c.rebus) l | None -> None) | _ -> None else None in @@ -333,13 +330,14 @@ let clenv_pose_metas_as_evars clenv dep_mvs = let ty = clenv_meta_type clenv mv in (* Postpone the evar-ization if dependent on another meta *) (* This assumes no cycle in the dependencies - is it correct ? *) - if occur_meta clenv.evd (EConstr.of_constr ty) then fold clenv (mvs@[mv]) + if occur_meta clenv.evd ty then fold clenv (mvs@[mv]) else let src = evar_source_of_meta mv clenv.evd in let src = adjust_meta_source clenv.evd mv src in let evd = Sigma.Unsafe.of_evar_map clenv.evd in - let Sigma (evar, evd, _) = new_evar (cl_env clenv) evd ~src (EConstr.of_constr ty) in + let Sigma (evar, evd, _) = new_evar (cl_env clenv) evd ~src ty in let evd = Sigma.to_evar_map evd in + let evar = EConstr.of_constr evar in let clenv = clenv_assign mv evar {clenv with evd=evd} in fold clenv mvs in fold clenv dep_mvs @@ -405,7 +403,7 @@ type arg_bindings = constr explicit_bindings * of cval, ctyp. *) let clenv_independent clenv = - let mvs = collect_metas clenv.evd (EConstr.of_constr (clenv_value clenv)) in + let mvs = collect_metas clenv.evd (clenv_value clenv) in let ctyp_mvs = (mk_freelisted (clenv_type clenv)).freemetas in let deps = Metaset.union (dependent_in_type_of_metas clenv mvs) ctyp_mvs in List.filter (fun mv -> not (Metaset.mem mv deps)) mvs @@ -483,15 +481,13 @@ let error_already_defined b = (str "Position " ++ int n ++ str" already defined.") let clenv_unify_binding_type clenv c t u = - let u = EConstr.of_constr u in - if isMeta (fst (decompose_appvect (whd_nored clenv.evd u))) then + if isMeta clenv.evd (EConstr.of_constr (fst (decompose_app_vect clenv.evd (EConstr.of_constr (whd_nored clenv.evd u))))) then (* Not enough information to know if some subtyping is needed *) CoerceToType, clenv, c else (* Enough information so as to try a coercion now *) try - let evd,c = w_coerce_to_type (cl_env clenv) clenv.evd (EConstr.of_constr c) (EConstr.of_constr t) u in - let c = EConstr.Unsafe.to_constr c in + let evd,c = w_coerce_to_type (cl_env clenv) clenv.evd c t u in TypeProcessed, { clenv with evd = evd }, c with | PretypeError (_,_,ActualTypeNotCoercible (_,_, @@ -501,9 +497,11 @@ let clenv_unify_binding_type clenv c t u = TypeNotProcessed, clenv, c let clenv_assign_binding clenv k c = - let k_typ = clenv_hnf_constr clenv (EConstr.of_constr (clenv_meta_type clenv k)) in - let c_typ = nf_betaiota clenv.evd (EConstr.of_constr (clenv_get_type_of clenv c)) in + let k_typ = clenv_hnf_constr clenv (clenv_meta_type clenv k) in + let c_typ = nf_betaiota clenv.evd (clenv_get_type_of clenv c) in + let c_typ = EConstr.of_constr c_typ in let status,clenv',c = clenv_unify_binding_type clenv c c_typ k_typ in + let c = EConstr.Unsafe.to_constr c in { clenv' with evd = meta_assign k (c,(Conv,status)) clenv'.evd } let clenv_match_args bl clenv = @@ -516,7 +514,7 @@ let clenv_match_args bl clenv = (fun clenv (loc,b,c) -> let k = meta_of_binder clenv loc mvs b in if meta_defined clenv.evd k then - if Term.eq_constr (fst (meta_fvalue clenv.evd k)).rebus c then clenv + if EConstr.eq_constr clenv.evd (EConstr.of_constr (fst (meta_fvalue clenv.evd k)).rebus) c then clenv else error_already_defined b else clenv_assign_binding clenv k c) @@ -525,7 +523,7 @@ let clenv_match_args bl clenv = exception NoSuchBinding let clenv_constrain_last_binding c clenv = - let all_mvs = collect_metas clenv.evd (EConstr.of_constr clenv.templval.rebus) in + let all_mvs = collect_metas clenv.evd clenv.templval.rebus in let k = try List.last all_mvs with Failure _ -> raise NoSuchBinding in clenv_assign_binding clenv k c @@ -560,8 +558,9 @@ let make_clenv_binding_gen hyps_only n env sigma (c,t) = function let clause = mk_clenv_from_env env sigma n (c,t) in clenv_constrain_dep_args hyps_only largs clause | ExplicitBindings lbind -> + let t = EConstr.of_constr (rename_bound_vars_as_displayed [] [] (EConstr.Unsafe.to_constr t)) in let clause = mk_clenv_from_env env sigma n - (c,rename_bound_vars_as_displayed [] [] t) + (c, t) in clenv_match_args lbind clause | NoBindings -> mk_clenv_from_env env sigma n (c,t) @@ -579,43 +578,49 @@ let make_clenv_binding env sigma = make_clenv_binding_gen false None env sigma (* Pretty-print *) let pr_clenv clenv = + let inj = EConstr.Unsafe.to_constr in h 0 - (str"TEMPL: " ++ print_constr clenv.templval.rebus ++ - str" : " ++ print_constr clenv.templtyp.rebus ++ fnl () ++ + (str"TEMPL: " ++ print_constr (inj clenv.templval.rebus) ++ + str" : " ++ print_constr (inj clenv.templtyp.rebus) ++ fnl () ++ pr_evar_map (Some 2) clenv.evd) (****************************************************************) (** Evar version of mk_clenv *) type hole = { - hole_evar : constr; - hole_type : types; + hole_evar : EConstr.constr; + hole_type : EConstr.types; hole_deps : bool; hole_name : Name.t; } type clause = { cl_holes : hole list; - cl_concl : types; + cl_concl : EConstr.types; } let make_evar_clause env sigma ?len t = + let open EConstr in + let open Vars in let bound = match len with | None -> -1 | Some n -> assert (0 <= n); n in (** FIXME: do the renaming online *) + let t = EConstr.Unsafe.to_constr t in let t = rename_bound_vars_as_displayed [] [] t in + let t = EConstr.of_constr t in let rec clrec (sigma, holes) n t = if n = 0 then (sigma, holes, t) - else match kind_of_term t with + else match EConstr.kind sigma t with | Cast (t, _, _) -> clrec (sigma, holes) n t | Prod (na, t1, t2) -> let store = Typeclasses.set_resolvable Evd.Store.empty false in let sigma = Sigma.Unsafe.of_evar_map sigma in - let Sigma (ev, sigma, _) = new_evar ~store env sigma (EConstr.of_constr t1) in + let Sigma (ev, sigma, _) = new_evar ~store env sigma t1 in let sigma = Sigma.to_evar_map sigma in - let dep = not (EConstr.Vars.noccurn sigma 1 (EConstr.of_constr t2)) in + let ev = EConstr.of_constr ev in + let dep = not (noccurn sigma 1 t2) in let hole = { hole_evar = ev; hole_type = t1; @@ -670,26 +675,28 @@ let evar_of_binder holes = function user_err (str "No such binder.") let define_with_type sigma env ev c = - let c = EConstr.of_constr c in - let t = Retyping.get_type_of env sigma (EConstr.of_constr ev) in + let open EConstr in + let t = Retyping.get_type_of env sigma ev in + let t = EConstr.of_constr t in let ty = Retyping.get_type_of env sigma c in let ty = EConstr.of_constr ty in let j = Environ.make_judge c ty in - let (sigma, j) = Coercion.inh_conv_coerce_to true (Loc.ghost) env sigma j (EConstr.of_constr t) in - let (ev, _) = destEvar ev in + let (sigma, j) = Coercion.inh_conv_coerce_to true (Loc.ghost) env sigma j t in + let (ev, _) = destEvar sigma ev in let sigma = Evd.define ev (EConstr.Unsafe.to_constr j.Environ.uj_val) sigma in sigma let solve_evar_clause env sigma hyp_only clause = function | NoBindings -> sigma | ImplicitBindings largs -> + let open EConstr in let fold holes h = if h.hole_deps then (** Some subsequent term uses the hole *) - let (ev, _) = destEvar h.hole_evar in - let is_dep hole = occur_evar sigma ev (EConstr.of_constr hole.hole_type) in + let (ev, _) = destEvar sigma h.hole_evar in + let is_dep hole = occur_evar sigma ev hole.hole_type in let in_hyp = List.exists is_dep holes in - let in_ccl = occur_evar sigma ev (EConstr.of_constr clause.cl_concl) in + let in_ccl = occur_evar sigma ev clause.cl_concl in let dep = if hyp_only then in_hyp && not in_ccl else in_hyp || in_ccl in let h = { h with hole_deps = dep } in h :: holes diff --git a/proofs/clenv.mli b/proofs/clenv.mli index e9236b1da..e4f6f9a91 100644 --- a/proofs/clenv.mli +++ b/proofs/clenv.mli @@ -14,6 +14,7 @@ open Names open Term open Environ open Evd +open EConstr open Unification open Misctypes @@ -28,8 +29,6 @@ type clausenv = { templtyp : constr freelisted (** its type *)} -val map_clenv : (constr -> constr) -> clausenv -> clausenv - (** subject of clenv (instantiated) *) val clenv_value : clausenv -> constr @@ -37,16 +36,16 @@ val clenv_value : clausenv -> constr val clenv_type : clausenv -> types (** substitute resolved metas *) -val clenv_nf_meta : clausenv -> constr -> constr +val clenv_nf_meta : clausenv -> Constr.constr -> Constr.constr (** type of a meta in clenv context *) val clenv_meta_type : clausenv -> metavariable -> types -val mk_clenv_from : Goal.goal sigma -> constr * types -> clausenv +val mk_clenv_from : Goal.goal sigma -> EConstr.constr * EConstr.types -> clausenv val mk_clenv_from_n : - Goal.goal sigma -> int option -> constr * types -> clausenv -val mk_clenv_type_of : Goal.goal sigma -> constr -> clausenv -val mk_clenv_from_env : env -> evar_map -> int option -> constr * types -> clausenv + Goal.goal sigma -> int option -> EConstr.constr * EConstr.types -> clausenv +val mk_clenv_type_of : Goal.goal sigma -> EConstr.constr -> clausenv +val mk_clenv_from_env : env -> evar_map -> int option -> EConstr.constr * EConstr.types -> clausenv (** Refresh the universes in a clenv *) val refresh_undefined_univs : clausenv -> clausenv * Univ.universe_level_subst @@ -92,18 +91,18 @@ val clenv_unify_meta_types : ?flags:unify_flags -> clausenv -> clausenv the optional int tells how many prods of the lemma have to be used use all of them if None *) val make_clenv_binding_env_apply : - env -> evar_map -> int option -> constr * constr -> constr bindings -> + env -> evar_map -> int option -> EConstr.constr * EConstr.constr -> constr bindings -> clausenv val make_clenv_binding_apply : - env -> evar_map -> int option -> constr * constr -> constr bindings -> + env -> evar_map -> int option -> EConstr.constr * EConstr.constr -> constr bindings -> clausenv val make_clenv_binding_env : - env -> evar_map -> constr * constr -> constr bindings -> clausenv + env -> evar_map -> EConstr.constr * EConstr.constr -> constr bindings -> clausenv val make_clenv_binding : - env -> evar_map -> constr * constr -> constr bindings -> clausenv + env -> evar_map -> EConstr.constr * EConstr.constr -> constr bindings -> clausenv (** if the clause is a product, add an extra meta for this product *) exception NotExtensibleClause @@ -134,9 +133,9 @@ val pr_clenv : clausenv -> Pp.std_ppcmds *) type hole = { - hole_evar : constr; + hole_evar : EConstr.constr; (** The hole itself. Guaranteed to be an evar. *) - hole_type : types; + hole_type : EConstr.types; (** Type of the hole in the current environment. *) hole_deps : bool; (** Whether the remainder of the clause was dependent in the hole. Note that @@ -148,10 +147,10 @@ type hole = { type clause = { cl_holes : hole list; - cl_concl : types; + cl_concl : EConstr.types; } -val make_evar_clause : env -> evar_map -> ?len:int -> types -> +val make_evar_clause : env -> evar_map -> ?len:int -> EConstr.types -> (evar_map * clause) (** An evar version of {!make_clenv_binding}. Given a type [t], [evar_environments env sigma ~len t bl] tries to eliminate at most [len] @@ -159,7 +158,7 @@ val make_evar_clause : env -> evar_map -> ?len:int -> types -> type together with the list of holes generated. Assumes that [t] is well-typed in the environment. *) -val solve_evar_clause : env -> evar_map -> bool -> clause -> constr bindings -> +val solve_evar_clause : env -> evar_map -> bool -> clause -> EConstr.constr bindings -> evar_map (** [solve_evar_clause env sigma hyps cl bl] tries to solve the holes contained in [cl] according to the [bl] argument. Assumes that [bl] are well-typed in diff --git a/proofs/clenvtac.ml b/proofs/clenvtac.ml index 5b9322bfe..539b1bcb2 100644 --- a/proofs/clenvtac.ml +++ b/proofs/clenvtac.ml @@ -11,6 +11,7 @@ open Names open Term open Termops open Evd +open EConstr open Refiner open Logic open Reduction @@ -26,19 +27,19 @@ open Proofview.Notations let clenv_cast_meta clenv = let rec crec u = - match kind_of_term u with + match EConstr.kind clenv.evd u with | App _ | Case _ -> crec_hd u - | Cast (c,_,_) when isMeta c -> u + | Cast (c,_,_) when isMeta clenv.evd c -> u | Proj (p, c) -> mkProj (p, crec_hd c) - | _ -> map_constr crec u + | _ -> EConstr.map clenv.evd crec u and crec_hd u = - match kind_of_term (strip_outer_cast clenv.evd (EConstr.of_constr u)) with + match EConstr.kind clenv.evd (EConstr.of_constr (strip_outer_cast clenv.evd u)) with | Meta mv -> (try let b = Typing.meta_type clenv.evd mv in - assert (not (occur_meta clenv.evd (EConstr.of_constr b))); - if occur_meta clenv.evd (EConstr.of_constr b) then u + assert (not (occur_meta clenv.evd b)); + if occur_meta clenv.evd b then u else mkCast (mkMeta mv, DEFAULTcast, b) with Not_found -> u) | App(f,args) -> mkApp (crec_hd f, Array.map crec args) @@ -95,7 +96,7 @@ let clenv_refine with_evars ?(with_classes=true) clenv = let clenv = { clenv with evd = evd' } in tclTHEN (tclEVARS (Evd.clear_metas evd')) - (refine_no_check (EConstr.of_constr (clenv_cast_meta clenv (clenv_value clenv)))) gl + (refine_no_check (clenv_cast_meta clenv (clenv_value clenv))) gl end open Unification @@ -143,7 +144,7 @@ let unify ?(flags=fail_quick_unif_flags) m = let n = Tacmach.New.pf_concl (Proofview.Goal.assume gl) in let evd = clear_metas (Tacmach.New.project gl) in try - let evd' = w_unify env evd CONV ~flags (EConstr.of_constr m) (EConstr.of_constr n) in + let evd' = w_unify env evd CONV ~flags m (EConstr.of_constr n) in Proofview.Unsafe.tclEVARSADVANCE evd' with e when CErrors.noncritical e -> Proofview.tclZERO e end } diff --git a/proofs/clenvtac.mli b/proofs/clenvtac.mli index 8a096b645..5b7164705 100644 --- a/proofs/clenvtac.mli +++ b/proofs/clenvtac.mli @@ -10,6 +10,7 @@ open Term open Clenv +open EConstr open Unification open Misctypes diff --git a/tactics/auto.ml b/tactics/auto.ml index 7462b8d85..2b654f563 100644 --- a/tactics/auto.ml +++ b/tactics/auto.ml @@ -84,11 +84,12 @@ let connect_hint_clenv poly (c, _, ctx) clenv gl = (** Refresh the instance of the hint *) let (subst, ctx) = Universes.fresh_universe_context_set_instance ctx in let map c = Vars.subst_univs_level_constr subst c in + let emap c = EConstr.Vars.subst_univs_level_constr subst c in let evd = Evd.merge_context_set Evd.univ_flexible evd ctx in (** Only metas are mentioning the old universes. *) let clenv = { - templval = Evd.map_fl map clenv.templval; - templtyp = Evd.map_fl map clenv.templtyp; + templval = Evd.map_fl emap clenv.templval; + templtyp = Evd.map_fl emap clenv.templtyp; evd = Evd.map_metas map evd; env = Proofview.Goal.env gl; } in diff --git a/tactics/autorewrite.ml b/tactics/autorewrite.ml index 80b9ec06e..b567344c9 100644 --- a/tactics/autorewrite.ml +++ b/tactics/autorewrite.ml @@ -257,12 +257,12 @@ type hypinfo = { let decompose_applied_relation metas env sigma c ctype left2right = let find_rel ty = - let eqclause = Clenv.mk_clenv_from_env env sigma None (c,ty) in + let eqclause = Clenv.mk_clenv_from_env env sigma None (EConstr.of_constr c,EConstr.of_constr ty) in let eqclause = if metas then eqclause else clenv_pose_metas_as_evars eqclause (Evd.undefined_metas eqclause.evd) in - let (equiv, args) = decompose_app (Clenv.clenv_type eqclause) in + let (equiv, args) = decompose_app (EConstr.Unsafe.to_constr (Clenv.clenv_type eqclause)) in let rec split_last_two = function | [c1;c2] -> [],(c1, c2) | x::y::z -> @@ -276,7 +276,7 @@ let decompose_applied_relation metas env sigma c ctype left2right = in (* if not (evd_convertible env eqclause.evd ty1 ty2) then None *) (* else *) - Some { hyp_cl=eqclause; hyp_prf=(Clenv.clenv_value eqclause); hyp_ty = ty; + Some { hyp_cl=eqclause; hyp_prf=EConstr.Unsafe.to_constr (Clenv.clenv_value eqclause); hyp_ty = ty; hyp_car=ty1; hyp_rel=mkApp (equiv, Array.of_list others); hyp_l2r=left2right; hyp_left=c1; hyp_right=c2; } with Not_found -> None diff --git a/tactics/class_tactics.ml b/tactics/class_tactics.ml index a2699ba8d..a8768b6ed 100644 --- a/tactics/class_tactics.ml +++ b/tactics/class_tactics.ml @@ -227,6 +227,7 @@ let e_give_exact flags poly (c,clenv) gl = else c, gl in let t1 = pf_unsafe_type_of gl (EConstr.of_constr c) in + let t1 = EConstr.of_constr t1 in Proofview.V82.of_tactic (Clenvtac.unify ~flags t1 <*> exact_no_check c) gl let unify_e_resolve poly flags = { enter = begin fun gls (c,_,clenv) -> @@ -247,6 +248,7 @@ let unify_resolve_refine poly flags = { enter = begin fun gls ((c, t, ctx),n,clenv) -> let env = Proofview.Goal.env gls in let concl = Proofview.Goal.concl gls in + let concl = EConstr.of_constr concl in Refine.refine ~unsafe:true { Sigma.run = fun sigma -> let sigma = Sigma.to_evar_map sigma in let sigma, term, ty = @@ -259,17 +261,20 @@ let unify_resolve_refine poly flags = let sigma = Evd.merge_context_set Evd.univ_flexible sigma ctx in sigma, c, t in + let open EConstr in + let ty = EConstr.of_constr ty in + let term = EConstr.of_constr term in let sigma', cl = Clenv.make_evar_clause env sigma ?len:n ty in - let term = applistc term (List.map (fun x -> x.hole_evar) cl.cl_holes) in + let term = applist (term, List.map (fun x -> x.hole_evar) cl.cl_holes) in let sigma' = let evdref = ref sigma' in if not (Evarconv.e_cumul env ~ts:flags.core_unify_flags.modulo_delta - evdref (EConstr.of_constr cl.cl_concl) (EConstr.of_constr concl)) then - Type_errors.error_actual_type env + evdref cl.cl_concl concl) then + Pretype_errors.error_actual_type_core env sigma' {Environ.uj_val = term; Environ.uj_type = cl.cl_concl} concl; !evdref - in Sigma.here (EConstr.of_constr term) (Sigma.Unsafe.of_evar_map sigma') } + in Sigma.here term (Sigma.Unsafe.of_evar_map sigma') } end } (** Dealing with goals of the form A -> B and hints of the form @@ -279,9 +284,11 @@ let clenv_of_prods poly nprods (c, clenv) gl = let (c, _, _) = c in if poly || Int.equal nprods 0 then Some (None, clenv) else + let c = EConstr.of_constr c in let sigma = Tacmach.New.project gl in - let ty = Retyping.get_type_of (Proofview.Goal.env gl) sigma (EConstr.of_constr c) in - let diff = nb_prod sigma (EConstr.of_constr ty) - nprods in + let ty = Retyping.get_type_of (Proofview.Goal.env gl) sigma c in + let ty = EConstr.of_constr ty in + let diff = nb_prod sigma ty - nprods in if Pervasives.(>=) diff 0 then (* Was Some clenv... *) Some (Some diff, @@ -1515,7 +1522,7 @@ let autoapply c i gl = let flags = auto_unif_flags Evar.Set.empty (Hints.Hint_db.transparent_state (Hints.searchtable_map i)) in let cty = pf_unsafe_type_of gl (EConstr.of_constr c) in - let ce = mk_clenv_from gl (c,cty) in + let ce = mk_clenv_from gl (EConstr.of_constr c,EConstr.of_constr cty) in let tac = { enter = fun gl -> (unify_e_resolve false flags).enter gl ((c,cty,Univ.ContextSet.empty),0,ce) } in Proofview.V82.of_tactic (Proofview.Goal.nf_enter tac) gl diff --git a/tactics/eauto.ml b/tactics/eauto.ml index 2fad4fcf7..7b07c9309 100644 --- a/tactics/eauto.ml +++ b/tactics/eauto.ml @@ -31,9 +31,10 @@ let eauto_unif_flags = auto_flags_of_state full_transparent_state let e_give_exact ?(flags=eauto_unif_flags) c = Proofview.Goal.enter { enter = begin fun gl -> let t1 = Tacmach.New.pf_unsafe_type_of gl (EConstr.of_constr c) in + let t1 = EConstr.of_constr t1 in let t2 = Tacmach.New.pf_concl (Proofview.Goal.assume gl) in let sigma = Tacmach.New.project gl in - if occur_existential sigma (EConstr.of_constr t1) || occur_existential sigma (EConstr.of_constr t2) then + if occur_existential sigma t1 || occur_existential sigma (EConstr.of_constr t2) then Tacticals.New.tclTHEN (Clenvtac.unify ~flags t1) (exact_no_check c) else exact_check c end } diff --git a/tactics/equality.ml b/tactics/equality.ml index ad80d2d1f..fbf461f6f 100644 --- a/tactics/equality.ml +++ b/tactics/equality.ml @@ -144,7 +144,7 @@ let freeze_initial_evars sigma flags clause = (* We take evars of the type: this may include old evars! For excluding *) (* all old evars, including the ones occurring in the rewriting lemma, *) (* we would have to take the clenv_value *) - let newevars = Evd.evars_of_term (clenv_type clause) in + let newevars = Evd.evars_of_term (EConstr.Unsafe.to_constr (clenv_type clause)) in let evars = fold_undefined (fun evk _ evars -> if Evar.Set.mem evk newevars then evars @@ -165,8 +165,11 @@ let side_tac tac sidetac = let instantiate_lemma_all frzevars gl c ty l l2r concl = let env = Proofview.Goal.env gl in + let c = EConstr.of_constr c in + let ty = EConstr.of_constr ty in + let l = Miscops.map_bindings EConstr.of_constr l in let eqclause = pf_apply Clenv.make_clenv_binding gl (c,ty) l in - let (equiv, args) = decompose_appvect (Clenv.clenv_type eqclause) in + let (equiv, args) = decompose_appvect (EConstr.Unsafe.to_constr (Clenv.clenv_type eqclause)) in let arglen = Array.length args in let () = if arglen < 2 then error "The term provided is not an applied relation." in let c1 = args.(arglen - 2) in @@ -181,8 +184,11 @@ let instantiate_lemma_all frzevars gl c ty l l2r concl = in List.map try_occ occs let instantiate_lemma gl c ty l l2r concl = - let sigma, ct = pf_type_of gl (EConstr.of_constr c) in + let c = EConstr.of_constr c in + let sigma, ct = pf_type_of gl c in let t = try snd (reduce_to_quantified_ind (pf_env gl) sigma (EConstr.of_constr ct)) with UserError _ -> ct in + let t = EConstr.of_constr t in + let l = Miscops.map_bindings EConstr.of_constr l in let eqclause = Clenv.make_clenv_binding (pf_env gl) sigma (c,t) l in [eqclause] @@ -975,9 +981,11 @@ let eq_baseid = Id.of_string "e" let apply_on_clause (f,t) clause = let sigma = clause.evd in + let f = EConstr.of_constr f in + let t = EConstr.of_constr t in let f_clause = mk_clenv_from_env clause.env sigma None (f,t) in let argmv = - (match kind_of_term (last_arg f_clause.evd (EConstr.of_constr f_clause.templval.Evd.rebus)) with + (match kind_of_term (last_arg f_clause.evd f_clause.templval.Evd.rebus) with | Meta mv -> mv | _ -> user_err (str "Ill-formed clause applicator.")) in clenv_fchain ~with_univs:false argmv f_clause clause @@ -992,7 +1000,6 @@ let discr_positions env sigma (lbeq,eqn,(t,t1,t2)) eq_clause cpath dirn = let pf_ty = mkArrow eqn absurd_term in let absurd_clause = apply_on_clause (pf,pf_ty) eq_clause in let pf = Clenvtac.clenv_value_cast_meta absurd_clause in - let pf = EConstr.of_constr pf in Proofview.Unsafe.tclEVARS sigma <*> Proofview.tclEFFECTS eff <*> tclTHENS (assert_after Anonymous absurd_term) @@ -1011,13 +1018,17 @@ let discrEq (lbeq,_,(t,t1,t2) as u) eq_clause = let onEquality with_evars tac (c,lbindc) = Proofview.Goal.nf_enter { enter = begin fun gl -> + let c = EConstr.of_constr c in + let lbindc = Miscops.map_bindings EConstr.of_constr lbindc in let type_of = pf_unsafe_type_of gl in let reduce_to_quantified_ind = pf_apply Tacred.reduce_to_quantified_ind gl in - let t = type_of (EConstr.of_constr c) in + let t = type_of c in let t' = try snd (reduce_to_quantified_ind (EConstr.of_constr t)) with UserError _ -> t in + let t' = EConstr.of_constr t' in let eq_clause = pf_apply make_clenv_binding gl (c,t') lbindc in let eq_clause' = Clenvtac.clenv_pose_dependent_evars with_evars eq_clause in let eqn = clenv_type eq_clause' in + let eqn = EConstr.Unsafe.to_constr eqn in let (eq,u,eq_args) = find_this_eq_data_decompose gl eqn in tclTHEN (Proofview.Unsafe.tclEVARS eq_clause'.evd) @@ -1371,7 +1382,8 @@ let inject_at_positions env sigma l2r (eq,_,(t,t1,t2)) eq_clause posns tac = let sigma, pf_typ = Typing.type_of env sigma (EConstr.of_constr pf) in let inj_clause = apply_on_clause (pf,pf_typ) eq_clause in let pf = Clenvtac.clenv_value_cast_meta inj_clause in - let ty = simplify_args env sigma (clenv_type inj_clause) in + let ty = simplify_args env sigma (EConstr.Unsafe.to_constr (clenv_type inj_clause)) in + let pf = EConstr.Unsafe.to_constr pf in evdref := sigma; Some (pf, ty) with Failure _ -> None @@ -1405,7 +1417,7 @@ let injEqThen tac l2r (eq,_,(t,t1,t2) as u) eq_clause = tclZEROMSG (str"Nothing to inject.") | Inr posns -> inject_at_positions env sigma l2r u eq_clause posns - (tac (clenv_value eq_clause)) + (tac (EConstr.Unsafe.to_constr (clenv_value eq_clause))) let get_previous_hyp_position id gl = let rec aux dest = function @@ -1464,10 +1476,10 @@ let decompEqThen ntac (lbeq,_,(t,t1,t2) as u) clause = | Inl (cpath, (_,dirn), _) -> discr_positions env sigma u clause cpath dirn | Inr [] -> (* Change: do not fail, simplify clear this trivial hyp *) - ntac (clenv_value clause) 0 + ntac (EConstr.Unsafe.to_constr (clenv_value clause)) 0 | Inr posns -> inject_at_positions env sigma true u clause posns - (ntac (clenv_value clause)) + (ntac (EConstr.Unsafe.to_constr (clenv_value clause))) end } let dEqThen with_evars ntac = function @@ -1478,9 +1490,11 @@ let dEq with_evars = dEqThen with_evars (fun clear_flag c x -> (apply_clear_request clear_flag (use_clear_hyp_by_default ()) c)) -let intro_decomp_eq tac data cl = +let intro_decomp_eq tac data (c, t) = Proofview.Goal.enter { enter = begin fun gl -> - let cl = pf_apply make_clenv_binding gl cl NoBindings in + let c = EConstr.of_constr c in + let t = EConstr.of_constr t in + let cl = pf_apply make_clenv_binding gl (c, t) NoBindings in decompEqThen (fun _ -> tac) data cl end } diff --git a/tactics/hints.ml b/tactics/hints.ml index 57358bb76..ea95fb1ad 100644 --- a/tactics/hints.ml +++ b/tactics/hints.ml @@ -276,9 +276,11 @@ let strip_params env c = let instantiate_hint env sigma p = let mk_clenv (c, cty, ctx) = let sigma = Evd.merge_context_set univ_flexible sigma ctx in + let c = EConstr.of_constr c in + let cty = EConstr.of_constr cty in let cl = mk_clenv_from_env env sigma None (c,cty) in {cl with templval = - { cl.templval with rebus = strip_params env cl.templval.rebus }; + { cl.templval with rebus = EConstr.of_constr (strip_params env (EConstr.Unsafe.to_constr cl.templval.rebus)) }; env = empty_env} in let code = match p.code.obj with @@ -765,9 +767,9 @@ let make_apply_entry env sigma (eapply,hnf,verbose) pri poly ?(name=PathAny) (c, match kind_of_term cty with | Prod _ -> let sigma' = Evd.merge_context_set univ_flexible sigma ctx in - let ce = mk_clenv_from_env env sigma' None (c,cty) in + let ce = mk_clenv_from_env env sigma' None (EConstr.of_constr c,EConstr.of_constr cty) in let c' = clenv_type (* ~reduce:false *) ce in - let pat = Patternops.pattern_of_constr env ce.evd (EConstr.of_constr c') in + let pat = Patternops.pattern_of_constr env ce.evd c' in let hd = try head_pattern_bound pat with BoundPattern -> failwith "make_apply_entry" in @@ -912,10 +914,10 @@ let make_trivial env sigma poly ?(name=PathAny) r = let sigma = Evd.merge_context_set univ_flexible sigma ctx in let t = hnf_constr env sigma (EConstr.of_constr (unsafe_type_of env sigma (EConstr.of_constr c))) in let hd = head_of_constr_reference sigma (EConstr.of_constr (head_constr sigma t)) in - let ce = mk_clenv_from_env env sigma None (c,t) in + let ce = mk_clenv_from_env env sigma None (EConstr.of_constr c,EConstr.of_constr t) in (Some hd, { pri=1; poly = poly; - pat = Some (Patternops.pattern_of_constr env ce.evd (EConstr.of_constr (clenv_type ce))); + pat = Some (Patternops.pattern_of_constr env ce.evd (clenv_type ce)); name = name; db = None; secvars = secvars_of_constr env c; diff --git a/tactics/leminv.ml b/tactics/leminv.ml index 85910355e..16a048af8 100644 --- a/tactics/leminv.ml +++ b/tactics/leminv.ml @@ -258,8 +258,8 @@ let add_inversion_lemma_exn na com comsort bool tac = let lemInv id c gls = try - let clause = mk_clenv_type_of gls c in - let clause = clenv_constrain_last_binding (mkVar id) clause in + let clause = mk_clenv_type_of gls (EConstr.of_constr c) in + let clause = clenv_constrain_last_binding (EConstr.mkVar id) clause in Proofview.V82.of_tactic (Clenvtac.res_pf clause ~flags:(Unification.elim_flags ()) ~with_evars:false) gls with | NoSuchBinding -> diff --git a/tactics/tacticals.ml b/tactics/tacticals.ml index 02909243d..459947051 100644 --- a/tactics/tacticals.ml +++ b/tactics/tacticals.ml @@ -622,20 +622,22 @@ module New = struct (* c should be of type A1->.. An->B with B an inductive definition *) let general_elim_then_using mk_elim isrec allnames tac predicate ind (c, t) = + let c = EConstr.of_constr c in + let t = EConstr.of_constr t in Proofview.Goal.nf_enter { enter = begin fun gl -> let sigma, elim = Tacmach.New.of_old (mk_elim ind) gl in Proofview.tclTHEN (Proofview.Unsafe.tclEVARS sigma) (Proofview.Goal.nf_enter { enter = begin fun gl -> let indclause = Tacmach.New.of_old (fun gl -> mk_clenv_from gl (c, t)) gl in (* applying elimination_scheme just a little modified *) - let elimclause = Tacmach.New.of_old (fun gls -> mk_clenv_from gls (elim,Tacmach.New.pf_unsafe_type_of gl (EConstr.of_constr elim))) gl in + let elimclause = Tacmach.New.of_old (fun gls -> mk_clenv_from gls (EConstr.of_constr elim,EConstr.of_constr (Tacmach.New.pf_unsafe_type_of gl (EConstr.of_constr elim)))) gl in let indmv = - match kind_of_term (last_arg elimclause.evd (EConstr.of_constr elimclause.templval.Evd.rebus)) with + match kind_of_term (last_arg elimclause.evd elimclause.templval.Evd.rebus) with | Meta mv -> mv | _ -> anomaly (str"elimination") in let pmv = - let p, _ = decompose_app elimclause.templtyp.Evd.rebus in + let p, _ = decompose_app (EConstr.Unsafe.to_constr elimclause.templtyp.Evd.rebus) in match kind_of_term p with | Meta p -> p | _ -> @@ -655,11 +657,11 @@ module New = struct let elimclause' = match predicate with | None -> elimclause' - | Some p -> clenv_unify ~flags Reduction.CONV (mkMeta pmv) p elimclause' + | Some p -> clenv_unify ~flags Reduction.CONV (EConstr.mkMeta pmv) (EConstr.of_constr p) elimclause' in let clenv' = Tacmach.New.of_old (clenv_unique_resolver ~flags elimclause') gl in let after_tac i = - let (hd,largs) = decompose_app clenv'.templtyp.Evd.rebus in + let (hd,largs) = decompose_app (EConstr.Unsafe.to_constr clenv'.templtyp.Evd.rebus) in let ba = { branchsign = branchsigns.(i); branchnames = brnames.(i); nassums = List.length branchsigns.(i); diff --git a/tactics/tactics.ml b/tactics/tactics.ml index b9a219a2c..f262aefa7 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -1301,11 +1301,11 @@ let clenv_refine_in ?(sidecond_first=false) with_evars ?(with_classes=true) else clenv in let new_hyp_typ = clenv_type clenv in + let new_hyp_typ = EConstr.Unsafe.to_constr new_hyp_typ in if not with_evars then check_unresolved_evars_of_metas sigma0 clenv; if not with_evars && occur_meta clenv.evd (EConstr.of_constr new_hyp_typ) then error_uninstantiated_metas new_hyp_typ clenv; let new_hyp_prf = clenv_value clenv in - let new_hyp_prf = EConstr.of_constr new_hyp_prf in let exact_tac = Proofview.V82.tactic (Tacmach.refine_no_check new_hyp_prf) in let naming = NamingMustBe (dloc,targetid) in let with_clear = do_replace (Some id) naming in @@ -1396,9 +1396,12 @@ let elimination_clause_scheme with_evars ?(with_classes=true) ?(flags=elim_flags let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in let elim = contract_letin_in_lam_header elim in + let bindings = Miscops.map_bindings EConstr.of_constr bindings in + let elim = EConstr.of_constr elim in + let elimty = EConstr.of_constr elimty in let elimclause = make_clenv_binding env sigma (elim, elimty) bindings in let indmv = - (match kind_of_term (nth_arg i elimclause.templval.rebus) with + (match kind_of_term (nth_arg i (EConstr.Unsafe.to_constr elimclause.templval.rebus)) with | Meta mv -> mv | _ -> user_err ~hdr:"elimination_clause" (str "The type of elimination clause is not well-formed.")) @@ -1438,8 +1441,10 @@ let general_elim with_evars clear_flag (c, lbindc) elim = let sigma = Tacmach.New.project gl in let ct = Retyping.get_type_of env sigma (EConstr.of_constr c) in let t = try snd (reduce_to_quantified_ind env sigma (EConstr.of_constr ct)) with UserError _ -> ct in + let t = EConstr.of_constr t in let elimtac = elimination_clause_scheme with_evars in - let indclause = make_clenv_binding env sigma (c, t) lbindc in + let lbindc = Miscops.map_bindings EConstr.of_constr lbindc in + let indclause = make_clenv_binding env sigma (EConstr.of_constr c, t) lbindc in let sigma = meta_merge sigma (clear_metas indclause.evd) in Proofview.Unsafe.tclEVARS sigma <*> Tacticals.New.tclTHEN @@ -1561,8 +1566,11 @@ let elimination_in_clause_scheme with_evars ?(flags=elim_flags ()) let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in let elim = contract_letin_in_lam_header elim in + let elim = EConstr.of_constr elim in + let elimty = EConstr.of_constr elimty in + let bindings = Miscops.map_bindings EConstr.of_constr bindings in let elimclause = make_clenv_binding env sigma (elim, elimty) bindings in - let indmv = destMeta (nth_arg i elimclause.templval.rebus) in + let indmv = destMeta (nth_arg i (EConstr.Unsafe.to_constr elimclause.templval.rebus)) in let hypmv = try match List.remove Int.equal indmv (clenv_independent elimclause) with | [a] -> a @@ -1570,12 +1578,13 @@ let elimination_in_clause_scheme with_evars ?(flags=elim_flags ()) with Failure _ -> user_err ~hdr:"elimination_clause" (str "The type of elimination clause is not well-formed.") in let elimclause' = clenv_fchain ~flags indmv elimclause indclause in - let hyp = mkVar id in - let hyp_typ = Retyping.get_type_of env sigma (EConstr.of_constr hyp) in + let hyp = EConstr.mkVar id in + let hyp_typ = Retyping.get_type_of env sigma hyp in + let hyp_typ = EConstr.of_constr hyp_typ in let hypclause = mk_clenv_from_env env sigma (Some 0) (hyp, hyp_typ) in let elimclause'' = clenv_fchain_in id ~flags hypmv elimclause' hypclause in let new_hyp_typ = clenv_type elimclause'' in - if Term.eq_constr hyp_typ new_hyp_typ then + if EConstr.eq_constr sigma hyp_typ new_hyp_typ then user_err ~hdr:"general_rewrite_in" (str "Nothing to rewrite in " ++ pr_id id ++ str"."); clenv_refine_in with_evars id id sigma elimclause'' @@ -1728,9 +1737,11 @@ let general_apply with_delta with_destruct with_evars clear_flag (loc,(c,lbind)) let thm_ty0 = nf_betaiota sigma (EConstr.of_constr (Retyping.get_type_of env sigma (EConstr.of_constr c))) in let try_apply thm_ty nprod = try - let n = nb_prod_modulo_zeta sigma (EConstr.of_constr thm_ty) - nprod in + let thm_ty = EConstr.of_constr thm_ty in + let n = nb_prod_modulo_zeta sigma thm_ty - nprod in if n<0 then error "Applied theorem has not enough premisses."; - let clause = make_clenv_binding_apply env sigma (Some n) (c,thm_ty) lbind in + let lbind = Miscops.map_bindings EConstr.of_constr lbind in + let clause = make_clenv_binding_apply env sigma (Some n) (EConstr.of_constr c,thm_ty) lbind in Clenvtac.res_pf clause ~with_evars ~flags with exn when catchable_exception exn -> Proofview.tclZERO exn @@ -1851,7 +1862,9 @@ let progress_with_clause flags innerclause clause = with Not_found -> error "Unable to unify." let apply_in_once_main flags innerclause env sigma (d,lbind) = - let thm = nf_betaiota sigma (EConstr.of_constr (Retyping.get_type_of env sigma (EConstr.of_constr d))) in + let d = EConstr.of_constr d in + let thm = nf_betaiota sigma (EConstr.of_constr (Retyping.get_type_of env sigma d)) in + let thm = EConstr.of_constr thm in let rec aux clause = try progress_with_clause flags innerclause clause with e when CErrors.noncritical e -> @@ -1859,6 +1872,7 @@ let apply_in_once_main flags innerclause env sigma (d,lbind) = try aux (clenv_push_prod clause) with NotExtensibleClause -> iraise e in + let lbind = Miscops.map_bindings EConstr.of_constr lbind in aux (make_clenv_binding env sigma (d,thm) lbind) let apply_in_once sidecond_first with_delta with_destruct with_evars naming @@ -1870,7 +1884,7 @@ let apply_in_once sidecond_first with_delta with_destruct with_evars naming let flags = if with_delta then default_unify_flags () else default_no_delta_unify_flags () in let t' = Tacmach.New.pf_get_hyp_typ id gl in - let innerclause = mk_clenv_from_env env sigma (Some 0) (mkVar id,t') in + let innerclause = mk_clenv_from_env env sigma (Some 0) (EConstr.mkVar id,EConstr.of_constr t') in let targetid = find_name true (LocalAssum (Anonymous,t')) naming gl in let rec aux idstoclear with_destruct c = Proofview.Goal.enter { enter = begin fun gl -> @@ -2939,10 +2953,12 @@ let specialize (c,lbind) ipat = let sigma = Typeclasses.resolve_typeclasses env sigma in sigma, nf_evar sigma c else - let clause = make_clenv_binding env sigma (c,Retyping.get_type_of env sigma (EConstr.of_constr c)) lbind in + let c = EConstr.of_constr c in + let lbind = Miscops.map_bindings EConstr.of_constr lbind in + let clause = make_clenv_binding env sigma (c,EConstr.of_constr (Retyping.get_type_of env sigma c)) lbind in let flags = { (default_unify_flags ()) with resolve_evars = true } in let clause = clenv_unify_meta_types ~flags clause in - let (thd,tstack) = whd_nored_stack clause.evd (EConstr.of_constr (clenv_value clause)) in + let (thd,tstack) = whd_nored_stack clause.evd (clenv_value clause) in let rec chk = function | [] -> [] | t::l -> if occur_meta clause.evd t then [] else EConstr.Unsafe.to_constr t :: chk l @@ -4107,7 +4123,7 @@ let get_eliminator elim dep s gl = of lid are parameters (first ones), the other are arguments. Returns the clause obtained. *) let recolle_clenv i params args elimclause gl = - let _,arr = destApp elimclause.templval.rebus in + let _,arr = destApp (EConstr.Unsafe.to_constr elimclause.templval.rebus) in let lindmv = Array.map (fun x -> @@ -4132,6 +4148,8 @@ let recolle_clenv i params args elimclause gl = (* from_n (Some 0) means that x should be taken "as is" without trying to unify (which would lead to trying to apply it to evars if y is a product). *) + let x = EConstr.of_constr x in + let y = EConstr.of_constr y in let indclause = Tacmach.New.of_old (fun gl -> mk_clenv_from_n gl (Some 0) (x,y)) gl in let elimclause' = clenv_fchain ~with_univs:false i acc indclause in elimclause') @@ -4149,6 +4167,9 @@ let induction_tac with_evars params indvars elim = (* elimclause contains this: (elimc ?i ?j ?k...?l) *) let elimc = contract_letin_in_lam_header elimc in let elimc = mkCast (elimc, DEFAULTcast, elimt) in + let elimc = EConstr.of_constr elimc in + let elimt = EConstr.of_constr elimt in + let lbindelimc = Miscops.map_bindings EConstr.of_constr lbindelimc in let elimclause = pf_apply make_clenv_binding gl (elimc,elimt) lbindelimc in (* elimclause' is built from elimclause by instanciating all args and params. *) let elimclause' = recolle_clenv i params indvars elimclause gl in @@ -4294,11 +4315,14 @@ let use_bindings env sigma elim must_be_closed (c,lbind) typ = typ in let rec find_clause typ = try + let typ = EConstr.of_constr typ in + let c = EConstr.of_constr c in + let lbind = Miscops.map_bindings EConstr.of_constr lbind in let indclause = make_clenv_binding env sigma (c,typ) lbind in - if must_be_closed && occur_meta indclause.evd (EConstr.of_constr (clenv_value indclause)) then + if must_be_closed && occur_meta indclause.evd (clenv_value indclause) then error "Need a fully applied argument."; (* We lose the possibility of coercions in with-bindings *) - let (sigma, c) = pose_all_metas_as_evars env indclause.evd (EConstr.of_constr (clenv_value indclause)) in + let (sigma, c) = pose_all_metas_as_evars env indclause.evd (clenv_value indclause) in Sigma.Unsafe.of_pair (EConstr.Unsafe.to_constr c, sigma) with e when catchable_exception e -> try find_clause (try_red_product env sigma (EConstr.of_constr typ)) @@ -4308,13 +4332,15 @@ let use_bindings env sigma elim must_be_closed (c,lbind) typ = let check_expected_type env sigma (elimc,bl) elimt = (* Compute the expected template type of the term in case a using clause is given *) - let sign,_ = splay_prod env sigma (EConstr.of_constr elimt) in + let open EConstr in + let elimt = EConstr.of_constr elimt in + let sign,_ = splay_prod env sigma elimt in let n = List.length sign in if n == 0 then error "Scheme cannot be applied."; let sigma,cl = make_evar_clause env sigma ~len:(n - 1) elimt in let sigma = solve_evar_clause env sigma true cl bl in - let (_,u,_) = destProd cl.cl_concl in - fun t -> Evarconv.e_cumul env (ref sigma) t (EConstr.of_constr u) + let (_,u,_) = destProd sigma cl.cl_concl in + fun t -> Evarconv.e_cumul env (ref sigma) t u let check_enough_applied env sigma elim = let sigma = Sigma.to_evar_map sigma in @@ -4327,6 +4353,7 @@ let check_enough_applied env sigma elim = | Some elimc -> let elimt = Retyping.get_type_of env sigma (EConstr.of_constr (fst elimc)) in let scheme = compute_elim_sig ~elimc elimt in + let elimc = Miscops.map_with_bindings EConstr.of_constr elimc in match scheme.indref with | None -> (* in the absence of information, do not assume it may be @@ -4607,12 +4634,13 @@ let simple_destruct = function let elim_scheme_type elim t = Proofview.Goal.nf_enter { enter = begin fun gl -> + let elim = EConstr.of_constr elim in let clause = Tacmach.New.of_old (fun gl -> mk_clenv_type_of gl elim) gl in - match kind_of_term (last_arg clause.templval.rebus) with + match kind_of_term (last_arg (EConstr.Unsafe.to_constr clause.templval.rebus)) with | Meta mv -> let clause' = (* t is inductive, then CUMUL or CONV is irrelevant *) - clenv_unify ~flags:(elim_flags ()) Reduction.CUMUL t + clenv_unify ~flags:(elim_flags ()) Reduction.CUMUL (EConstr.of_constr t) (clenv_meta_type clause mv) clause in Clenvtac.res_pf clause' ~flags:(elim_flags ()) ~with_evars:false | _ -> anomaly (Pp.str "elim_scheme_type") -- cgit v1.2.3 From 45562afa065aadc207dca4e904e309d835cb66ef Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sat, 12 Nov 2016 01:28:45 +0100 Subject: Tacticals API using EConstr. --- engine/termops.ml | 2 +- engine/termops.mli | 2 +- proofs/clenv.ml | 2 +- proofs/clenv.mli | 2 +- tactics/elim.ml | 6 +++--- tactics/equality.ml | 2 +- tactics/inv.ml | 6 ++++-- tactics/tacticals.ml | 20 ++++++++++---------- tactics/tacticals.mli | 14 +++++++------- 9 files changed, 29 insertions(+), 27 deletions(-) diff --git a/engine/termops.ml b/engine/termops.ml index 83f07d2c6..5581b1656 100644 --- a/engine/termops.ml +++ b/engine/termops.ml @@ -268,7 +268,7 @@ let rec drop_extra_implicit_args sigma c = match EConstr.kind sigma c with (* Get the last arg of an application *) let last_arg sigma c = match EConstr.kind sigma c with - | App (f,cl) -> EConstr.Unsafe.to_constr (Array.last cl) + | App (f,cl) -> Array.last cl | _ -> anomaly (Pp.str "last_arg") (* Get the last arg of an application *) diff --git a/engine/termops.mli b/engine/termops.mli index 27b3ea53c..02e8dfeae 100644 --- a/engine/termops.mli +++ b/engine/termops.mli @@ -201,7 +201,7 @@ val nb_prod : Evd.evar_map -> EConstr.t -> int val nb_prod_modulo_zeta : Evd.evar_map -> EConstr.t -> int (** Get the last arg of a constr intended to be an application *) -val last_arg : Evd.evar_map -> EConstr.t -> constr +val last_arg : Evd.evar_map -> EConstr.t -> EConstr.constr (** Force the decomposition of a term as an applicative one *) val decompose_app_vect : Evd.evar_map -> EConstr.t -> constr * constr array diff --git a/proofs/clenv.ml b/proofs/clenv.ml index 3e3889cbb..fd88e3c51 100644 --- a/proofs/clenv.ml +++ b/proofs/clenv.ml @@ -44,7 +44,7 @@ type clausenv = { let cl_env ce = ce.env let cl_sigma ce = ce.evd -let clenv_nf_meta clenv c = nf_meta clenv.evd c +let clenv_nf_meta clenv c = EConstr.of_constr (nf_meta clenv.evd (EConstr.Unsafe.to_constr c)) let clenv_term clenv c = meta_instance clenv.evd c let clenv_meta_type clenv mv = Typing.meta_type clenv.evd mv let clenv_value clenv = meta_instance clenv.evd clenv.templval diff --git a/proofs/clenv.mli b/proofs/clenv.mli index e4f6f9a91..f7ff4bed3 100644 --- a/proofs/clenv.mli +++ b/proofs/clenv.mli @@ -36,7 +36,7 @@ val clenv_value : clausenv -> constr val clenv_type : clausenv -> types (** substitute resolved metas *) -val clenv_nf_meta : clausenv -> Constr.constr -> Constr.constr +val clenv_nf_meta : clausenv -> EConstr.constr -> EConstr.constr (** type of a meta in clenv context *) val clenv_meta_type : clausenv -> metavariable -> types diff --git a/tactics/elim.ml b/tactics/elim.ml index bcb1c05cc..fe36085b8 100644 --- a/tactics/elim.ml +++ b/tactics/elim.ml @@ -55,7 +55,7 @@ Another example : *) let elimHypThen tac id = - elimination_then tac (mkVar id) + elimination_then tac (EConstr.mkVar id) let rec general_decompose_on_hyp recognizer = ifOnHyp recognizer (general_decompose_aux recognizer) (fun _ -> Proofview.tclUNIT()) @@ -125,7 +125,7 @@ let h_decompose_and = decompose_and (* The tactic Double performs a double induction *) let simple_elimination c = - elimination_then (fun _ -> tclIDTAC) c + elimination_then (fun _ -> tclIDTAC) (EConstr.of_constr c) let induction_trailer abs_i abs_j bargs = tclTHEN @@ -166,7 +166,7 @@ let double_ind h1 h2 = (onLastHypId (fun id -> elimination_then - (introElimAssumsThen (induction_trailer abs_i abs_j)) (mkVar id)))) + (introElimAssumsThen (induction_trailer abs_i abs_j)) (EConstr.mkVar id)))) end } let h_double_induction = double_ind diff --git a/tactics/equality.ml b/tactics/equality.ml index fbf461f6f..fa4164bb9 100644 --- a/tactics/equality.ml +++ b/tactics/equality.ml @@ -985,7 +985,7 @@ let apply_on_clause (f,t) clause = let t = EConstr.of_constr t in let f_clause = mk_clenv_from_env clause.env sigma None (f,t) in let argmv = - (match kind_of_term (last_arg f_clause.evd f_clause.templval.Evd.rebus) with + (match EConstr.kind sigma (last_arg f_clause.evd f_clause.templval.Evd.rebus) with | Meta mv -> mv | _ -> user_err (str "Ill-formed clause applicator.")) in clenv_fchain ~with_univs:false argmv f_clause clause diff --git a/tactics/inv.ml b/tactics/inv.ml index 2f5186f81..60f1c3542 100644 --- a/tactics/inv.ml +++ b/tactics/inv.ml @@ -443,7 +443,8 @@ let raw_inversion inv_kind id status names = let msg = str "The type of " ++ pr_id id ++ str " is not inductive." in CErrors.user_err msg in - let IndType (indf,realargs) = find_rectype env sigma (EConstr.of_constr t) in + let t = EConstr.of_constr t in + let IndType (indf,realargs) = find_rectype env sigma t in let evdref = ref sigma in let (elim_predicate, args) = make_inv_predicate env evdref indf realargs id status concl in @@ -463,13 +464,14 @@ let raw_inversion inv_kind id status names = in let neqns = List.length realargs in let as_mode = names != None in + let elim_predicate = EConstr.of_constr elim_predicate in let tac = (tclTHENS (assert_before Anonymous cut_concl) [case_tac names (introCaseAssumsThen false (* ApplyOn not supported by inversion *) (rewrite_equations_tac as_mode inv_kind id neqns)) - (Some elim_predicate) ind (c, t); + (Some elim_predicate) ind (EConstr.of_constr c,t); onLastHypId (fun id -> tclTHEN (refined id) reflexivity)]) in Sigma.Unsafe.of_pair (tac, sigma) diff --git a/tactics/tacticals.ml b/tactics/tacticals.ml index 459947051..0546132c1 100644 --- a/tactics/tacticals.ml +++ b/tactics/tacticals.ml @@ -147,9 +147,9 @@ let ifOnHyp pred tac1 tac2 id gl = type branch_args = { ity : pinductive; (* the type we were eliminating on *) - largs : constr list; (* its arguments *) + largs : EConstr.constr list; (* its arguments *) branchnum : int; (* the branch number *) - pred : constr; (* the predicate we used *) + pred : EConstr.constr; (* the predicate we used *) nassums : int; (* number of assumptions/letin to be introduced *) branchsign : bool list; (* the signature of the branch. true=assumption, false=let-in *) @@ -622,8 +622,7 @@ module New = struct (* c should be of type A1->.. An->B with B an inductive definition *) let general_elim_then_using mk_elim isrec allnames tac predicate ind (c, t) = - let c = EConstr.of_constr c in - let t = EConstr.of_constr t in + let open EConstr in Proofview.Goal.nf_enter { enter = begin fun gl -> let sigma, elim = Tacmach.New.of_old (mk_elim ind) gl in Proofview.tclTHEN (Proofview.Unsafe.tclEVARS sigma) @@ -632,13 +631,13 @@ module New = struct (* applying elimination_scheme just a little modified *) let elimclause = Tacmach.New.of_old (fun gls -> mk_clenv_from gls (EConstr.of_constr elim,EConstr.of_constr (Tacmach.New.pf_unsafe_type_of gl (EConstr.of_constr elim)))) gl in let indmv = - match kind_of_term (last_arg elimclause.evd elimclause.templval.Evd.rebus) with + match EConstr.kind elimclause.evd (last_arg elimclause.evd elimclause.templval.Evd.rebus) with | Meta mv -> mv | _ -> anomaly (str"elimination") in let pmv = - let p, _ = decompose_app (EConstr.Unsafe.to_constr elimclause.templtyp.Evd.rebus) in - match kind_of_term p with + let p, _ = decompose_app elimclause.evd elimclause.templtyp.Evd.rebus in + match EConstr.kind elimclause.evd p with | Meta p -> p | _ -> let name_elim = @@ -657,11 +656,11 @@ module New = struct let elimclause' = match predicate with | None -> elimclause' - | Some p -> clenv_unify ~flags Reduction.CONV (EConstr.mkMeta pmv) (EConstr.of_constr p) elimclause' + | Some p -> clenv_unify ~flags Reduction.CONV (EConstr.mkMeta pmv) p elimclause' in let clenv' = Tacmach.New.of_old (clenv_unique_resolver ~flags elimclause') gl in let after_tac i = - let (hd,largs) = decompose_app (EConstr.Unsafe.to_constr clenv'.templtyp.Evd.rebus) in + let (hd,largs) = decompose_app clenv'.evd clenv'.templtyp.Evd.rebus in let ba = { branchsign = branchsigns.(i); branchnames = brnames.(i); nassums = List.length branchsigns.(i); @@ -680,7 +679,8 @@ module New = struct let elimination_then tac c = Proofview.Goal.nf_enter { enter = begin fun gl -> - let (ind,t) = pf_reduce_to_quantified_ind gl (EConstr.of_constr (pf_unsafe_type_of gl (EConstr.of_constr c))) in + let (ind,t) = pf_reduce_to_quantified_ind gl (EConstr.of_constr (pf_unsafe_type_of gl c)) in + let t = EConstr.of_constr t in let isrec,mkelim = match (Global.lookup_mind (fst (fst ind))).mind_record with | None -> true,gl_make_elim diff --git a/tactics/tacticals.mli b/tactics/tacticals.mli index 18cf03c51..974bf83a3 100644 --- a/tactics/tacticals.mli +++ b/tactics/tacticals.mli @@ -97,17 +97,17 @@ val onClauseLR : (Id.t option -> tactic) -> clause -> tactic (** {6 Elimination tacticals. } *) -type branch_args = { +type branch_args = private { ity : pinductive; (** the type we were eliminating on *) - largs : constr list; (** its arguments *) + largs : EConstr.constr list; (** its arguments *) branchnum : int; (** the branch number *) - pred : constr; (** the predicate we used *) + pred : EConstr.constr; (** the predicate we used *) nassums : int; (** number of assumptions/letin to be introduced *) branchsign : bool list; (** the signature of the branch. true=assumption, false=let-in *) branchnames : intro_patterns} -type branch_assumptions = { +type branch_assumptions = private { ba : branch_args; (** the branch args *) assums : Context.Named.t} (** the list of assumptions introduced *) @@ -253,15 +253,15 @@ module New : sig val elimination_then : (branch_args -> unit Proofview.tactic) -> - constr -> unit Proofview.tactic + EConstr.constr -> unit Proofview.tactic val case_then_using : or_and_intro_pattern option -> (branch_args -> unit Proofview.tactic) -> - constr option -> pinductive -> Term.constr * Term.types -> unit Proofview.tactic + EConstr.constr option -> pinductive -> EConstr.constr * EConstr.types -> unit Proofview.tactic val case_nodep_then_using : or_and_intro_pattern option -> (branch_args -> unit Proofview.tactic) -> - constr option -> pinductive -> Term.constr * Term.types -> unit Proofview.tactic + EConstr.constr option -> pinductive -> EConstr.constr * EConstr.types -> unit Proofview.tactic val elim_on_ba : (branch_assumptions -> unit Proofview.tactic) -> branch_args -> unit Proofview.tactic val case_on_ba : (branch_assumptions -> unit Proofview.tactic) -> branch_args -> unit Proofview.tactic -- cgit v1.2.3 From 771be16883c8c47828f278ce49545716918764c4 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sat, 12 Nov 2016 01:52:15 +0100 Subject: Hipattern API using EConstr. --- engine/termops.ml | 9 +++ engine/termops.mli | 2 + ltac/tauto.ml | 11 ++- plugins/firstorder/formula.ml | 17 ++--- plugins/funind/invfun.ml | 4 +- tactics/contradiction.ml | 10 +-- tactics/elim.ml | 4 +- tactics/eqdecide.ml | 12 ++-- tactics/equality.ml | 72 +++++++++++++------ tactics/equality.mli | 4 +- tactics/hipattern.ml | 156 ++++++++++++++++++++++-------------------- tactics/hipattern.mli | 24 ++++--- tactics/inv.ml | 2 +- tactics/tactics.ml | 38 +++++++--- tactics/tactics.mli | 2 +- toplevel/indschemes.ml | 2 +- 16 files changed, 224 insertions(+), 145 deletions(-) diff --git a/engine/termops.ml b/engine/termops.ml index 5581b1656..59dbb73f5 100644 --- a/engine/termops.ml +++ b/engine/termops.ml @@ -829,6 +829,15 @@ let global_of_constr sigma c = | Var id -> VarRef id, Univ.Instance.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 diff --git a/engine/termops.mli b/engine/termops.mli index 02e8dfeae..abc9caa98 100644 --- a/engine/termops.mli +++ b/engine/termops.mli @@ -268,6 +268,8 @@ val is_section_variable : Id.t -> bool val global_of_constr : Evd.evar_map -> EConstr.constr -> Globnames.global_reference puniverses +val is_global : Evd.evar_map -> Globnames.global_reference -> EConstr.constr -> bool + val isGlobalRef : Evd.evar_map -> EConstr.t -> bool val is_template_polymorphic : env -> Evd.evar_map -> EConstr.t -> bool diff --git a/ltac/tauto.ml b/ltac/tauto.ml index 6eab003b5..11996af73 100644 --- a/ltac/tauto.ml +++ b/ltac/tauto.ml @@ -113,7 +113,7 @@ let split = Tactics.split_with_bindings false [Misctypes.NoBindings] let is_empty _ ist = Proofview.tclEVARMAP >>= fun sigma -> - if is_empty_type sigma (assoc_var "X1" ist) then idtac else fail + if is_empty_type sigma (EConstr.of_constr (assoc_var "X1" ist)) then idtac else fail (* Strictly speaking, this exceeds the propositional fragment as it matches also equality types (and solves them if a reflexivity) *) @@ -121,9 +121,10 @@ let is_unit_or_eq _ ist = Proofview.tclEVARMAP >>= fun sigma -> let flags = assoc_flags ist in let test = if flags.strict_unit then is_unit_type else is_unit_or_eq_type in - if test sigma (assoc_var "X1" ist) then idtac else fail + if test sigma (EConstr.of_constr (assoc_var "X1" ist)) then idtac else fail let bugged_is_binary t = + let t = EConstr.Unsafe.to_constr t in isApp t && let (hdapp,args) = decompose_app t in match (kind_of_term hdapp) with @@ -138,6 +139,7 @@ let is_conj _ ist = Proofview.tclEVARMAP >>= fun sigma -> let flags = assoc_flags ist in let ind = assoc_var "X1" ist in + let ind = EConstr.of_constr ind in if (not flags.binary_mode_bugged_detection || bugged_is_binary ind) && is_conjunction sigma ~strict:flags.strict_in_hyp_and_ccl @@ -149,6 +151,7 @@ let flatten_contravariant_conj _ ist = Proofview.tclEVARMAP >>= fun sigma -> let flags = assoc_flags ist in let typ = assoc_var "X1" ist in + let typ = EConstr.of_constr typ in let c = assoc_var "X2" ist in let hyp = assoc_var "id" ist in match match_with_conjunction sigma @@ -156,6 +159,7 @@ let flatten_contravariant_conj _ ist = ~onlybinary:flags.binary_mode typ with | Some (_,args) -> + let args = List.map EConstr.Unsafe.to_constr args in let newtyp = List.fold_right mkArrow args c in let intros = tclMAP (fun _ -> intro) args in let by = tclTHENLIST [intros; apply hyp; split; assumption] in @@ -168,6 +172,7 @@ let is_disj _ ist = Proofview.tclEVARMAP >>= fun sigma -> let flags = assoc_flags ist in let t = assoc_var "X1" ist in + let t = EConstr.of_constr t in if (not flags.binary_mode_bugged_detection || bugged_is_binary t) && is_disjunction sigma ~strict:flags.strict_in_hyp_and_ccl @@ -179,6 +184,7 @@ let flatten_contravariant_disj _ ist = Proofview.tclEVARMAP >>= fun sigma -> let flags = assoc_flags ist in let typ = assoc_var "X1" ist in + let typ = EConstr.of_constr typ in let c = assoc_var "X2" ist in let hyp = assoc_var "id" ist in match match_with_disjunction sigma @@ -186,6 +192,7 @@ let flatten_contravariant_disj _ ist = ~onlybinary:flags.binary_mode typ with | Some (_,args) -> + let args = List.map EConstr.Unsafe.to_constr args in let map i arg = let typ = mkArrow arg c in let ci = Tactics.constructor_tac false None (succ i) Misctypes.NoBindings in diff --git a/plugins/firstorder/formula.ml b/plugins/firstorder/formula.ml index 60e9196af..96b991e1f 100644 --- a/plugins/firstorder/formula.ml +++ b/plugins/firstorder/formula.ml @@ -79,15 +79,16 @@ type kind_of_formula= let kind_of_formula gl term = let normalize=special_nf gl in let cciterm=special_whd gl term in - match match_with_imp_term (project gl) cciterm with - Some (a,b)-> Arrow(a,(pop (EConstr.of_constr b))) + match match_with_imp_term (project gl) (EConstr.of_constr cciterm) with + Some (a,b)-> Arrow(EConstr.Unsafe.to_constr a,(pop b)) |_-> - match match_with_forall_term (project gl) cciterm with - Some (_,a,b)-> Forall(a,b) + match match_with_forall_term (project gl) (EConstr.of_constr cciterm) with + Some (_,a,b)-> Forall(EConstr.Unsafe.to_constr a,EConstr.Unsafe.to_constr b) |_-> - match match_with_nodep_ind (project gl) cciterm with + match match_with_nodep_ind (project gl) (EConstr.of_constr cciterm) with Some (i,l,n)-> - let ind,u=destInd i in + let l = List.map EConstr.Unsafe.to_constr l in + let ind,u=EConstr.destInd (project gl) i in let (mib,mip) = Global.lookup_inductive ind in let nconstr=Array.length mip.mind_consnames in if Int.equal nconstr 0 then @@ -108,8 +109,8 @@ let kind_of_formula gl term = else Or((ind,u),l,is_trivial) | _ -> - match match_with_sigma_type (project gl) cciterm with - Some (i,l)-> Exists((destInd i),l) + match match_with_sigma_type (project gl) (EConstr.of_constr cciterm) with + Some (i,l)-> Exists((EConstr.destInd (project gl) i),List.map EConstr.Unsafe.to_constr l) |_-> Atom (normalize cciterm) type atoms = {positive:constr list;negative:constr list} diff --git a/plugins/funind/invfun.ml b/plugins/funind/invfun.ml index 635925562..b2419b1a5 100644 --- a/plugins/funind/invfun.ml +++ b/plugins/funind/invfun.ml @@ -598,9 +598,9 @@ let rec reflexivity_with_destruct_cases g = | Some id -> match kind_of_term (pf_unsafe_type_of g (EConstr.mkVar id)) with | App(eq,[|_;t1;t2|]) when eq_constr eq eq_ind -> - if Equality.discriminable (pf_env g) (project g) t1 t2 + if Equality.discriminable (pf_env g) (project g) (EConstr.of_constr t1) (EConstr.of_constr t2) then Proofview.V82.of_tactic (Equality.discrHyp id) g - else if Equality.injectable (pf_env g) (project g) t1 t2 + else if Equality.injectable (pf_env g) (project g) (EConstr.of_constr t1) (EConstr.of_constr t2) then tclTHENSEQ [Proofview.V82.of_tactic (Equality.injHyp None id);thin [id];intros_with_rewrite] g else tclIDTAC g | _ -> tclIDTAC g diff --git a/tactics/contradiction.ml b/tactics/contradiction.ml index 2058b95a6..a8be704b2 100644 --- a/tactics/contradiction.ml +++ b/tactics/contradiction.ml @@ -66,12 +66,12 @@ let contradiction_context = let id = NamedDecl.get_id d in let typ = nf_evar sigma (NamedDecl.get_type d) in let typ = whd_all env sigma (EConstr.of_constr typ) in - if is_empty_type sigma typ then + if is_empty_type sigma (EConstr.of_constr typ) then simplest_elim (mkVar id) else match kind_of_term typ with - | Prod (na,t,u) when is_empty_type sigma u -> + | Prod (na,t,u) when is_empty_type sigma (EConstr.of_constr u) -> let is_unit_or_eq = - if use_negated_unit_or_eq_type () then match_with_unit_or_eq_type sigma t + if use_negated_unit_or_eq_type () then match_with_unit_or_eq_type sigma (EConstr.of_constr t) else None in Tacticals.New.tclORELSE (match is_unit_or_eq with @@ -105,7 +105,7 @@ let is_negation_of env sigma typ t = match kind_of_term (whd_all env sigma t) with | Prod (na,t,u) -> let u = nf_evar sigma u in - is_empty_type sigma u && is_conv_leq env sigma (EConstr.of_constr typ) (EConstr.of_constr t) + is_empty_type sigma (EConstr.of_constr u) && is_conv_leq env sigma (EConstr.of_constr typ) (EConstr.of_constr t) | _ -> false let contradiction_term (c,lbind as cl) = @@ -115,7 +115,7 @@ let contradiction_term (c,lbind as cl) = let type_of = Tacmach.New.pf_unsafe_type_of gl in let typ = type_of (EConstr.of_constr c) in let _, ccl = splay_prod env sigma (EConstr.of_constr typ) in - if is_empty_type sigma ccl then + if is_empty_type sigma (EConstr.of_constr ccl) then Tacticals.New.tclTHEN (elim false None cl None) (Tacticals.New.tclTRY assumption) diff --git a/tactics/elim.ml b/tactics/elim.ml index fe36085b8..d00e504ff 100644 --- a/tactics/elim.ml +++ b/tactics/elim.ml @@ -108,12 +108,12 @@ let decompose_these c l = let decompose_and c = general_decompose - (fun sigma (_,t) -> is_record sigma t) + (fun sigma (_,t) -> is_record sigma (EConstr.of_constr t)) c let decompose_or c = general_decompose - (fun sigma (_,t) -> is_disjunction sigma t) + (fun sigma (_,t) -> is_disjunction sigma (EConstr.of_constr t)) c let h_decompose l c = decompose_these c l diff --git a/tactics/eqdecide.ml b/tactics/eqdecide.ml index d1b14a907..ed81d748a 100644 --- a/tactics/eqdecide.ml +++ b/tactics/eqdecide.ml @@ -141,8 +141,8 @@ open Proofview.Notations (* spiwack: a small wrapper around [Hipattern]. *) -let match_eqdec c = - try Proofview.tclUNIT (match_eqdec c) +let match_eqdec sigma c = + try Proofview.tclUNIT (match_eqdec sigma c) with PatternMatchingFailure -> Proofview.tclZERO PatternMatchingFailure (* /spiwack *) @@ -171,7 +171,9 @@ let solveEqBranch rectype = begin Proofview.Goal.enter { enter = begin fun gl -> let concl = pf_nf_concl gl in - match_eqdec concl >>= fun (eqonleft,op,lhs,rhs,_) -> + let concl = EConstr.of_constr concl in + let sigma = project gl in + match_eqdec sigma concl >>= fun (eqonleft,op,lhs,rhs,_) -> let (mib,mip) = Global.lookup_inductive rectype in let nparams = mib.mind_nparams in let getargs l = List.skipn nparams (snd (decompose_app l)) in @@ -196,7 +198,9 @@ let decideGralEquality = begin Proofview.Goal.enter { enter = begin fun gl -> let concl = pf_nf_concl gl in - match_eqdec concl >>= fun (eqonleft,_,c1,c2,typ) -> + let concl = EConstr.of_constr concl in + let sigma = project gl in + match_eqdec sigma concl >>= fun (eqonleft,_,c1,c2,typ) -> let headtyp = hd_app (pf_compute gl (EConstr.of_constr typ)) in begin match kind_of_term headtyp with | Ind (mi,_) -> Proofview.tclUNIT mi diff --git a/tactics/equality.ml b/tactics/equality.ml index fa4164bb9..e1a8d2bdb 100644 --- a/tactics/equality.ml +++ b/tactics/equality.ml @@ -448,8 +448,9 @@ let general_rewrite_ebindings_clause cls lft2rgt occs frzevars dep_proof_ok ?tac let env = Proofview.Goal.env gl in let ctype = get_type_of env sigma (EConstr.of_constr c) in let rels, t = decompose_prod_assum (whd_betaiotazeta sigma (EConstr.of_constr ctype)) in - match match_with_equality_type sigma t with + match match_with_equality_type sigma (EConstr.of_constr t) with | Some (hdcncl,args) -> (* Fast path: direct leibniz-like rewrite *) + let hdcncl = EConstr.Unsafe.to_constr hdcncl in let lft2rgt = adjust_rewriting_direction args lft2rgt in leibniz_rewrite_ebindings_clause cls lft2rgt tac c (it_mkProd_or_LetIn t rels) l with_evars frzevars dep_proof_ok hdcncl @@ -464,8 +465,9 @@ let general_rewrite_ebindings_clause cls lft2rgt occs frzevars dep_proof_ok ?tac Proofview.tclEVARMAP >>= fun sigma -> let env' = push_rel_context rels env in let rels',t' = splay_prod_assum env' sigma (EConstr.of_constr t) in (* Search for underlying eq *) - match match_with_equality_type sigma t' with + match match_with_equality_type sigma (EConstr.of_constr t') with | Some (hdcncl,args) -> + let hdcncl = EConstr.Unsafe.to_constr hdcncl in let lft2rgt = adjust_rewriting_direction args lft2rgt in leibniz_rewrite_ebindings_clause cls lft2rgt tac c (it_mkProd_or_LetIn t' (rels' @ rels)) l with_evars frzevars dep_proof_ok hdcncl @@ -768,7 +770,7 @@ let find_positions env sigma t1 t2 = let sorts = if !keep_proof_equalities_for_injection then [InSet;InType;InProp] else [InSet;InType] in - Inr (findrec sorts [] (EConstr.of_constr t1) (EConstr.of_constr t2)) + Inr (findrec sorts [] t1 t2) with DiscrFound (path,c1,c2) -> Inl (path,c1,c2) @@ -943,7 +945,7 @@ let gen_absurdity id = Proofview.Goal.enter { enter = begin fun gl -> let sigma = project gl in let hyp_typ = pf_get_hyp_typ id (Proofview.Goal.assume gl) in - let hyp_typ = pf_nf_evar gl hyp_typ in + let hyp_typ = EConstr.of_constr hyp_typ in if is_empty_type sigma hyp_typ then simplest_elim (mkVar id) @@ -991,6 +993,9 @@ let apply_on_clause (f,t) clause = clenv_fchain ~with_univs:false argmv f_clause clause let discr_positions env sigma (lbeq,eqn,(t,t1,t2)) eq_clause cpath dirn = + let t = EConstr.Unsafe.to_constr t in + let t1 = EConstr.Unsafe.to_constr t1 in + let t2 = EConstr.Unsafe.to_constr t2 in let e = next_ident_away eq_baseid (ids_of_context env) in let e_env = push_named (Context.Named.Declaration.LocalAssum (e,t)) env in let discriminator = @@ -1029,7 +1034,7 @@ let onEquality with_evars tac (c,lbindc) = let eq_clause' = Clenvtac.clenv_pose_dependent_evars with_evars eq_clause in let eqn = clenv_type eq_clause' in let eqn = EConstr.Unsafe.to_constr eqn in - let (eq,u,eq_args) = find_this_eq_data_decompose gl eqn in + let (eq,u,eq_args) = find_this_eq_data_decompose gl (EConstr.of_constr eqn) in tclTHEN (Proofview.Unsafe.tclEVARS eq_clause'.evd) (tac (eq,eqn,eq_args) eq_clause') @@ -1041,7 +1046,7 @@ let onNegatedEquality with_evars tac = let ccl = Proofview.Goal.concl gl in let env = Proofview.Goal.env gl in match kind_of_term (hnf_constr env sigma (EConstr.of_constr ccl)) with - | Prod (_,t,u) when is_empty_type sigma u -> + | Prod (_,t,u) when is_empty_type sigma (EConstr.of_constr u) -> tclTHEN introf (onLastHypId (fun id -> onEquality with_evars tac (mkVar id,NoBindings))) @@ -1320,6 +1325,7 @@ let inject_if_homogenous_dependent_pair ty = try let sigma = Tacmach.New.project gl in let eq,u,(t,t1,t2) = find_this_eq_data_decompose gl ty in + let t = EConstr.Unsafe.to_constr t in (* fetch the informations of the pair *) let ceq = Universes.constr_of_global Coqlib.glob_eq in let sigTconstr () = (Coqlib.build_sigma_type()).Coqlib.typ in @@ -1327,8 +1333,8 @@ let inject_if_homogenous_dependent_pair ty = (* check whether the equality deals with dep pairs or not *) let eqTypeDest = fst (decompose_app t) in if not (Globnames.is_global (sigTconstr()) eqTypeDest) then raise Exit; - let hd1,ar1 = decompose_app_vect sigma (EConstr.of_constr t1) and - hd2,ar2 = decompose_app_vect sigma (EConstr.of_constr t2) in + let hd1,ar1 = decompose_app_vect sigma t1 and + hd2,ar2 = decompose_app_vect sigma t2 in if not (Globnames.is_global (existTconstr()) hd1) then raise Exit; if not (Globnames.is_global (existTconstr()) hd2) then raise Exit; let ind,_ = try pf_apply find_mrectype gl (EConstr.of_constr ar1.(0)) with Not_found -> raise Exit in @@ -1369,6 +1375,9 @@ let simplify_args env sigma t = | _ -> t let inject_at_positions env sigma l2r (eq,_,(t,t1,t2)) eq_clause posns tac = + let t = EConstr.Unsafe.to_constr t in + let t1 = EConstr.Unsafe.to_constr t1 in + let t2 = EConstr.Unsafe.to_constr t2 in let e = next_ident_away eq_baseid (ids_of_context env) in let e_env = push_named (LocalAssum (e,t)) env in let evdref = ref sigma in @@ -1396,7 +1405,7 @@ let inject_at_positions env sigma l2r (eq,_,(t,t1,t2)) eq_clause posns tac = (Tacticals.New.tclTHENFIRST (Proofview.tclIGNORE (Proofview.Monad.List.map (fun (pf,ty) -> tclTHENS (cut ty) - [inject_if_homogenous_dependent_pair ty; + [inject_if_homogenous_dependent_pair (EConstr.of_constr ty); Proofview.V82.tactic (Tacmach.refine (EConstr.of_constr pf))]) (if l2r then List.rev injectors else injectors))) (tac (List.length injectors))) @@ -1536,7 +1545,12 @@ let decomp_tuple_term env sigma c t = let rec decomprec inner_code ex exty = let iterated_decomp = try - let ({proj1=p1; proj2=p2}),(i,a,p,car,cdr) = find_sigma_data_decompose ex in + let ex = EConstr.of_constr ex in + let ({proj1=p1; proj2=p2}),(i,a,p,car,cdr) = find_sigma_data_decompose sigma ex in + let a = EConstr.Unsafe.to_constr a in + let p = EConstr.Unsafe.to_constr p in + let car = EConstr.Unsafe.to_constr car in + let cdr = EConstr.Unsafe.to_constr cdr in let car_code = applist (mkConstU (destConstRef p1,i),[a;p;inner_code]) and cdr_code = applist (mkConstU (destConstRef p2,i),[a;p;inner_code]) in let cdrtyp = beta_applist sigma (EConstr.of_constr p,[EConstr.of_constr car]) in @@ -1547,6 +1561,8 @@ let decomp_tuple_term env sigma c t = in decomprec (mkRel 1) c t let subst_tuple_term env sigma dep_pair1 dep_pair2 b = + let dep_pair1 = EConstr.Unsafe.to_constr dep_pair1 in + let dep_pair2 = EConstr.Unsafe.to_constr dep_pair2 in let sigma = Sigma.to_evar_map sigma in let typ = get_type_of env sigma (EConstr.of_constr dep_pair1) in (* We find all possible decompositions *) @@ -1583,7 +1599,7 @@ let cutSubstInConcl l2r eqn = Proofview.Goal.nf_s_enter { s_enter = begin fun gl -> let env = Proofview.Goal.env gl in let sigma = Proofview.Goal.sigma gl in - let (lbeq,u,(t,e1,e2)) = find_eq_data_decompose gl eqn in + let (lbeq,u,(t,e1,e2)) = find_eq_data_decompose gl (EConstr.of_constr eqn) in let typ = pf_concl gl in let (e1,e2) = if l2r then (e1,e2) else (e2,e1) in let Sigma ((typ, expected), sigma, p) = subst_tuple_term env sigma e1 e2 typ in @@ -1602,7 +1618,7 @@ let cutSubstInHyp l2r eqn id = Proofview.Goal.nf_s_enter { s_enter = begin fun gl -> let env = Proofview.Goal.env gl in let sigma = Proofview.Goal.sigma gl in - let (lbeq,u,(t,e1,e2)) = find_eq_data_decompose gl eqn in + let (lbeq,u,(t,e1,e2)) = find_eq_data_decompose gl (EConstr.of_constr eqn) in let typ = pf_get_hyp_typ id gl in let (e1,e2) = if l2r then (e1,e2) else (e2,e1) in let Sigma ((typ, expected), sigma, p) = subst_tuple_term env sigma e1 e2 typ in @@ -1682,20 +1698,21 @@ let restrict_to_eq_and_identity eq = (* compatibility *) not (is_global glob_identity eq) then raise Constr_matching.PatternMatchingFailure -exception FoundHyp of (Id.t * constr * bool) +exception FoundHyp of (Id.t * EConstr.constr * bool) (* tests whether hyp [c] is [x = t] or [t = x], [x] not occurring in [t] *) let is_eq_x gl x d = let id = NamedDecl.get_id d in try - let is_var id c = match kind_of_term c with + let is_var id c = match EConstr.kind (project gl) c with | Var id' -> Id.equal id id' | _ -> false in let c = pf_nf_evar gl (NamedDecl.get_type d) in + let c = EConstr.of_constr c in let (_,lhs,rhs) = pi3 (find_eq_data_decompose gl c) in - if (is_var x lhs) && not (local_occur_var (project gl) x (EConstr.of_constr rhs)) then raise (FoundHyp (id,rhs,true)); - if (is_var x rhs) && not (local_occur_var (project gl) x (EConstr.of_constr lhs)) then raise (FoundHyp (id,lhs,false)) + if (is_var x lhs) && not (local_occur_var (project gl) x rhs) then raise (FoundHyp (id,rhs,true)); + if (is_var x rhs) && not (local_occur_var (project gl) x lhs) then raise (FoundHyp (id,lhs,false)) with Constr_matching.PatternMatchingFailure -> () @@ -1753,7 +1770,7 @@ let subst_one_var dep_proof_ok x = user_err ~hdr:"Subst" (str "Cannot find any non-recursive equality over " ++ pr_id x ++ str".") - with FoundHyp res -> res in + with FoundHyp (id, c, b) -> (id, EConstr.Unsafe.to_constr c, b) in subst_one dep_proof_ok x res end } @@ -1788,7 +1805,9 @@ let subst_all ?(flags=default_subst_tactic_flags ()) () = let find_eq_data_decompose = find_eq_data_decompose gl in let select_equation_name decl = try - let lbeq,u,(_,x,y) = find_eq_data_decompose (NamedDecl.get_type decl) in + let lbeq,u,(_,x,y) = find_eq_data_decompose (EConstr.of_constr (NamedDecl.get_type decl)) in + let x = EConstr.Unsafe.to_constr x in + let y = EConstr.Unsafe.to_constr y in let eq = Universes.constr_of_global_univ (lbeq.eq,u) in if flags.only_leibniz then restrict_to_eq_and_identity eq; match kind_of_term x, kind_of_term y with @@ -1812,7 +1831,10 @@ let subst_all ?(flags=default_subst_tactic_flags ()) () = let env = Proofview.Goal.env gl in let find_eq_data_decompose = find_eq_data_decompose gl in let c = pf_get_hyp hyp gl |> NamedDecl.get_type in + let c = EConstr.of_constr c in let _,_,(_,x,y) = find_eq_data_decompose c in + let x = EConstr.Unsafe.to_constr x in + let y = EConstr.Unsafe.to_constr y in (* J.F.: added to prevent failure on goal containing x=x as an hyp *) if Term.eq_constr x y then Proofview.tclUNIT () else match kind_of_term x, kind_of_term y with @@ -1838,7 +1860,10 @@ let subst_all ?(flags=default_subst_tactic_flags ()) () = let find_eq_data_decompose = find_eq_data_decompose gl in let test (_,c) = try + let c = EConstr.of_constr c in let lbeq,u,(_,x,y) = find_eq_data_decompose c in + let x = EConstr.Unsafe.to_constr x in + let y = EConstr.Unsafe.to_constr y in let eq = Universes.constr_of_global_univ (lbeq.eq,u) in if flags.only_leibniz then restrict_to_eq_and_identity eq; (* J.F.: added to prevent failure on goal containing x=x as an hyp *) @@ -1858,21 +1883,24 @@ let subst_all ?(flags=default_subst_tactic_flags ()) () = let cond_eq_term_left c t gl = try + let t = EConstr.of_constr t in let (_,x,_) = pi3 (find_eq_data_decompose gl t) in - if pf_conv_x gl (EConstr.of_constr c) (EConstr.of_constr x) then true else failwith "not convertible" + if pf_conv_x gl (EConstr.of_constr c) x then true else failwith "not convertible" with Constr_matching.PatternMatchingFailure -> failwith "not an equality" let cond_eq_term_right c t gl = try + let t = EConstr.of_constr t in let (_,_,x) = pi3 (find_eq_data_decompose gl t) in - if pf_conv_x gl (EConstr.of_constr c) (EConstr.of_constr x) then false else failwith "not convertible" + if pf_conv_x gl (EConstr.of_constr c) x then false else failwith "not convertible" with Constr_matching.PatternMatchingFailure -> failwith "not an equality" let cond_eq_term c t gl = try + let t = EConstr.of_constr t in let (_,x,y) = pi3 (find_eq_data_decompose gl t) in - if pf_conv_x gl (EConstr.of_constr c) (EConstr.of_constr x) then true - else if pf_conv_x gl (EConstr.of_constr c) (EConstr.of_constr y) then false + if pf_conv_x gl (EConstr.of_constr c) x then true + else if pf_conv_x gl (EConstr.of_constr c) y then false else failwith "not convertible" with Constr_matching.PatternMatchingFailure -> failwith "not an equality" diff --git a/tactics/equality.mli b/tactics/equality.mli index 6a4a8126e..779d1e9b2 100644 --- a/tactics/equality.mli +++ b/tactics/equality.mli @@ -96,8 +96,8 @@ val cutRewriteInConcl : bool -> constr -> unit Proofview.tactic val rewriteInHyp : bool -> constr -> Id.t -> unit Proofview.tactic val rewriteInConcl : bool -> constr -> unit Proofview.tactic -val discriminable : env -> evar_map -> constr -> constr -> bool -val injectable : env -> evar_map -> constr -> constr -> bool +val discriminable : env -> evar_map -> EConstr.constr -> EConstr.constr -> bool +val injectable : env -> evar_map -> EConstr.constr -> EConstr.constr -> bool (* Subst *) diff --git a/tactics/hipattern.ml b/tactics/hipattern.ml index 5d78fd585..6681e5e49 100644 --- a/tactics/hipattern.ml +++ b/tactics/hipattern.ml @@ -12,6 +12,7 @@ open Util open Names open Term open Termops +open EConstr open Inductiveops open Constr_matching open Coqlib @@ -31,9 +32,9 @@ module RelDecl = Context.Rel.Declaration -- Eduardo (6/8/97). *) -type 'a matching_function = Evd.evar_map -> constr -> 'a option +type 'a matching_function = Evd.evar_map -> EConstr.constr -> 'a option -type testing_function = Evd.evar_map -> constr -> bool +type testing_function = Evd.evar_map -> EConstr.constr -> bool let mkmeta n = Nameops.make_ident "X" (Some n) let meta1 = mkmeta 1 @@ -44,10 +45,10 @@ let meta4 = mkmeta 4 let op2bool = function Some _ -> true | None -> false let match_with_non_recursive_type sigma t = - match kind_of_term t with + match EConstr.kind sigma t with | App _ -> - let (hdapp,args) = decompose_app t in - (match kind_of_term hdapp with + let (hdapp,args) = decompose_app sigma t in + (match EConstr.kind sigma hdapp with | Ind (ind,u) -> if (Global.lookup_mind (fst ind)).mind_finite == Decl_kinds.CoFinite then Some (hdapp,args) @@ -64,9 +65,9 @@ let is_non_recursive_type sigma t = op2bool (match_with_non_recursive_type sigma since they may appear in types of inductive constructors (see #2629) *) let rec has_nodep_prod_after n sigma c = - match kind_of_term c with + match EConstr.kind sigma c with | Prod (_,_,b) | LetIn (_,_,_,b) -> - ( n>0 || EConstr.Vars.noccurn sigma 1 (EConstr.of_constr b)) + ( n>0 || Vars.noccurn sigma 1 b) && (has_nodep_prod_after (n-1) sigma b) | _ -> true @@ -87,9 +88,11 @@ let is_lax_conjunction = function | Some false -> true | _ -> false +let prod_assum sigma t = fst (decompose_prod_assum sigma t) + let match_with_one_constructor sigma style onlybinary allow_rec t = - let (hdapp,args) = decompose_app t in - let res = match kind_of_term hdapp with + let (hdapp,args) = decompose_app sigma t in + let res = match EConstr.kind sigma hdapp with | Ind ind -> let (mib,mip) = Global.lookup_inductive (fst ind) in if Int.equal (Array.length mip.mind_consnames) 1 @@ -98,22 +101,23 @@ let match_with_one_constructor sigma style onlybinary allow_rec t = then if is_strict_conjunction style (* strict conjunction *) then let ctx = - (prod_assum (snd - (decompose_prod_n_assum mib.mind_nparams mip.mind_nf_lc.(0)))) in + (prod_assum sigma (snd + (decompose_prod_n_assum sigma mib.mind_nparams (EConstr.of_constr mip.mind_nf_lc.(0))))) in if List.for_all (fun decl -> let c = RelDecl.get_type decl in is_local_assum decl && - isRel c && - Int.equal (destRel c) mib.mind_nparams) ctx + Term.isRel c && + Int.equal (Term.destRel c) mib.mind_nparams) ctx then Some (hdapp,args) else None else - let ctyp = Term.prod_applist mip.mind_nf_lc.(0) args in - let cargs = List.map RelDecl.get_type (prod_assum ctyp) in + let ctyp = Termops.prod_applist sigma (EConstr.of_constr mip.mind_nf_lc.(0)) args in + let cargs = List.map RelDecl.get_type (prod_assum sigma ctyp) in if not (is_lax_conjunction style) || has_nodep_prod sigma ctyp then (* Record or non strict conjunction *) + let cargs = List.map EConstr.of_constr cargs in Some (hdapp,List.rev cargs) else None @@ -140,7 +144,7 @@ let is_record sigma t = let match_with_tuple sigma t = let t = match_with_one_constructor sigma None false true t in Option.map (fun (hd,l) -> - let ind = destInd hd in + let ind = destInd sigma hd in let (mib,mip) = Global.lookup_pinductive ind in let isrec = mis_is_recursive (fst ind,mib,mip) in (hd,l,isrec)) t @@ -154,14 +158,15 @@ let is_tuple sigma t = "Inductive I A1 ... An := C1 (_:A1) | ... | Cn : (_:An)" *) let test_strict_disjunction n lc = + let open Term in Array.for_all_i (fun i c -> match (prod_assum (snd (decompose_prod_n_assum n c))) with | [LocalAssum (_,c)] -> isRel c && Int.equal (destRel c) (n - i) | _ -> false) 0 lc let match_with_disjunction ?(strict=false) ?(onlybinary=false) sigma t = - let (hdapp,args) = decompose_app t in - let res = match kind_of_term hdapp with + let (hdapp,args) = decompose_app sigma t in + let res = match EConstr.kind sigma hdapp with | Ind (ind,u) -> let car = constructors_nrealargs ind in let (mib,mip) = Global.lookup_inductive ind in @@ -176,7 +181,7 @@ let match_with_disjunction ?(strict=false) ?(onlybinary=false) sigma t = None else let cargs = - Array.map (fun ar -> pi2 (destProd (Term.prod_applist ar args))) + Array.map (fun ar -> pi2 (destProd sigma (prod_applist sigma (EConstr.of_constr ar) args))) mip.mind_nf_lc in Some (hdapp,Array.to_list cargs) else @@ -194,8 +199,8 @@ let is_disjunction ?(strict=false) ?(onlybinary=false) sigma t = constructors *) let match_with_empty_type sigma t = - let (hdapp,args) = decompose_app t in - match (kind_of_term hdapp) with + let (hdapp,args) = decompose_app sigma t in + match EConstr.kind sigma hdapp with | Ind ind -> let (mib,mip) = Global.lookup_pinductive ind in let nconstr = Array.length mip.mind_consnames in @@ -208,8 +213,8 @@ let is_empty_type sigma t = op2bool (match_with_empty_type sigma t) Parameters and indices are allowed *) let match_with_unit_or_eq_type sigma t = - let (hdapp,args) = decompose_app t in - match (kind_of_term hdapp) with + let (hdapp,args) = decompose_app sigma t in + match EConstr.kind sigma hdapp with | Ind ind -> let (mib,mip) = Global.lookup_pinductive ind in let constr_types = mip.mind_nf_lc in @@ -276,13 +281,13 @@ let coq_refl_jm_pattern = open Globnames -let is_matching x y = is_matching (Global.env ()) Evd.empty x (EConstr.of_constr y) -let matches x y = matches (Global.env ()) Evd.empty x (EConstr.of_constr y) +let is_matching sigma x y = is_matching (Global.env ()) sigma x y +let matches sigma x y = matches (Global.env ()) sigma x y -let match_with_equation t = - if not (isApp t) then raise NoEquationFound; - let (hdapp,args) = destApp t in - match kind_of_term hdapp with +let match_with_equation sigma t = + if not (isApp sigma t) then raise NoEquationFound; + let (hdapp,args) = destApp sigma t in + match EConstr.kind sigma hdapp with | Ind (ind,u) -> if eq_gr (IndRef ind) glob_eq then Some (build_coq_eq_data()),hdapp, @@ -298,11 +303,11 @@ let match_with_equation t = let constr_types = mip.mind_nf_lc in let nconstr = Array.length mip.mind_consnames in if Int.equal nconstr 1 then - if is_matching coq_refl_leibniz1_pattern constr_types.(0) then + if is_matching sigma coq_refl_leibniz1_pattern (EConstr.of_constr constr_types.(0)) then None, hdapp, MonomorphicLeibnizEq(args.(0),args.(1)) - else if is_matching coq_refl_leibniz2_pattern constr_types.(0) then + else if is_matching sigma coq_refl_leibniz2_pattern (EConstr.of_constr constr_types.(0)) then None, hdapp, PolymorphicLeibnizEq(args.(0),args.(1),args.(2)) - else if is_matching coq_refl_jm_pattern constr_types.(0) then + else if is_matching sigma coq_refl_jm_pattern (EConstr.of_constr constr_types.(0)) then None, hdapp, HeterogenousEq(args.(0),args.(1),args.(2),args.(3)) else raise NoEquationFound else raise NoEquationFound @@ -319,8 +324,8 @@ let is_inductive_equality ind = Int.equal nconstr 1 && Int.equal (constructor_nrealargs (ind,1)) 0 let match_with_equality_type sigma t = - let (hdapp,args) = decompose_app t in - match (kind_of_term hdapp) with + let (hdapp,args) = decompose_app sigma t in + match EConstr.kind sigma hdapp with | Ind (ind,_) when is_inductive_equality ind -> Some (hdapp,args) | _ -> None @@ -331,23 +336,25 @@ let is_equality_type sigma t = op2bool (match_with_equality_type sigma t) (** X1 -> X2 **) let coq_arrow_pattern = mkPattern (mkGArrow (mkGPatVar "X1") (mkGPatVar "X2")) -let match_arrow_pattern t = - let result = matches coq_arrow_pattern t in +let match_arrow_pattern sigma t = + let result = matches sigma coq_arrow_pattern t in match Id.Map.bindings result with | [(m1,arg);(m2,mind)] -> assert (Id.equal m1 meta1 && Id.equal m2 meta2); (arg, mind) | _ -> anomaly (Pp.str "Incorrect pattern matching") let match_with_imp_term sigma c = - match kind_of_term c with - | Prod (_,a,b) when EConstr.Vars.noccurn sigma 1 (EConstr.of_constr b) -> Some (a,b) + match EConstr.kind sigma c with + | Prod (_,a,b) when Vars.noccurn sigma 1 b -> Some (a,b) | _ -> None let is_imp_term sigma c = op2bool (match_with_imp_term sigma c) let match_with_nottype sigma t = try - let (arg,mind) = match_arrow_pattern t in + let (arg,mind) = match_arrow_pattern sigma t in + let arg = EConstr.of_constr arg in + let mind = EConstr.of_constr mind in if is_empty_type sigma mind then Some (mind,arg) else None with PatternMatchingFailure -> None @@ -356,19 +363,19 @@ let is_nottype sigma t = op2bool (match_with_nottype sigma t) (* Forall *) let match_with_forall_term sigma c= - match kind_of_term c with + match EConstr.kind sigma c with | Prod (nam,a,b) -> Some (nam,a,b) | _ -> None let is_forall_term sigma c = op2bool (match_with_forall_term sigma c) let match_with_nodep_ind sigma t = - let (hdapp,args) = decompose_app t in - match (kind_of_term hdapp) with + let (hdapp,args) = decompose_app sigma t in + match EConstr.kind sigma hdapp with | Ind ind -> let (mib,mip) = Global.lookup_pinductive ind in if Array.length (mib.mind_packets)>1 then None else - let nodep_constr = has_nodep_prod_after mib.mind_nparams sigma in + let nodep_constr c = has_nodep_prod_after mib.mind_nparams sigma (EConstr.of_constr c) in if Array.for_all nodep_constr mip.mind_nf_lc then let params= if Int.equal mip.mind_nrealargs 0 then args else @@ -381,14 +388,14 @@ let match_with_nodep_ind sigma t = let is_nodep_ind sigma t = op2bool (match_with_nodep_ind sigma t) let match_with_sigma_type sigma t = - let (hdapp,args) = decompose_app t in - match (kind_of_term hdapp) with + let (hdapp,args) = decompose_app sigma t in + match EConstr.kind sigma hdapp with | Ind ind -> let (mib,mip) = Global.lookup_pinductive ind in if Int.equal (Array.length (mib.mind_packets)) 1 && (Int.equal mip.mind_nrealargs 0) && (Int.equal (Array.length mip.mind_consnames)1) && - has_nodep_prod_after (mib.mind_nparams+1) sigma mip.mind_nf_lc.(0) then + has_nodep_prod_after (mib.mind_nparams+1) sigma (EConstr.of_constr mip.mind_nf_lc.(0)) then (*allowing only 1 existential*) Some (hdapp,args) else @@ -408,17 +415,17 @@ let rec first_match matcher = function (*** Equality *) -let match_eq eqn (ref, hetero) = +let match_eq sigma eqn (ref, hetero) = let ref = try Lazy.force ref with e when CErrors.noncritical e -> raise PatternMatchingFailure in - match kind_of_term eqn with + match EConstr.kind sigma eqn with | App (c, [|t; x; y|]) -> - if not hetero && is_global ref c then PolymorphicLeibnizEq (t, x, y) + if not hetero && Termops.is_global sigma ref c then PolymorphicLeibnizEq (t, x, y) else raise PatternMatchingFailure | App (c, [|t; x; t'; x'|]) -> - if hetero && is_global ref c then HeterogenousEq (t, x, t', x') + if hetero && Termops.is_global sigma ref c then HeterogenousEq (t, x, t', x') else raise PatternMatchingFailure | _ -> raise PatternMatchingFailure @@ -430,27 +437,27 @@ let equalities = (coq_jmeq_ref, true), check_jmeq_loaded, build_coq_jmeq_data; (coq_identity_ref, false), no_check, build_coq_identity_data] -let find_eq_data eqn = (* fails with PatternMatchingFailure *) - let d,k = first_match (match_eq eqn) equalities in - let hd,u = destInd (fst (destApp eqn)) in +let find_eq_data sigma eqn = (* fails with PatternMatchingFailure *) + let d,k = first_match (match_eq sigma eqn) equalities in + let hd,u = destInd sigma (fst (destApp sigma eqn)) in d,u,k let extract_eq_args gl = function | MonomorphicLeibnizEq (e1,e2) -> - let t = pf_unsafe_type_of gl (EConstr.of_constr e1) in (t,e1,e2) + let t = pf_unsafe_type_of gl e1 in (EConstr.of_constr t,e1,e2) | PolymorphicLeibnizEq (t,e1,e2) -> (t,e1,e2) | HeterogenousEq (t1,e1,t2,e2) -> - if pf_conv_x gl (EConstr.of_constr t1) (EConstr.of_constr t2) then (t1,e1,e2) + if pf_conv_x gl t1 t2 then (t1,e1,e2) else raise PatternMatchingFailure let find_eq_data_decompose gl eqn = - let (lbeq,u,eq_args) = find_eq_data eqn in + let (lbeq,u,eq_args) = find_eq_data (project gl) eqn in (lbeq,u,extract_eq_args gl eq_args) let find_this_eq_data_decompose gl eqn = let (lbeq,u,eq_args) = try (*first_match (match_eq eqn) inversible_equalities*) - find_eq_data eqn + find_eq_data (project gl) eqn with PatternMatchingFailure -> user_err (str "No primitive equality found.") in let eq_args = @@ -463,7 +470,6 @@ let match_eq_nf gls eqn (ref, hetero) = let n = if hetero then 4 else 3 in let args = List.init n (fun i -> mkGPatVar ("X" ^ string_of_int (i + 1))) in let pat = mkPattern (mkGAppRef ref args) in - let eqn = EConstr.of_constr eqn in match Id.Map.bindings (pf_matches gls pat eqn) with | [(m1,t);(m2,x);(m3,y)] -> assert (Id.equal m1 meta1 && Id.equal m2 meta2 && Id.equal m3 meta3); @@ -478,12 +484,12 @@ let dest_nf_eq gls eqn = (*** Sigma-types *) -let match_sigma ex = - match kind_of_term ex with - | App (f, [| a; p; car; cdr |]) when is_global (Lazy.force coq_exist_ref) f -> - build_sigma (), (snd (destConstruct f), a, p, car, cdr) - | App (f, [| a; p; car; cdr |]) when is_global (Lazy.force coq_existT_ref) f -> - build_sigma_type (), (snd (destConstruct f), a, p, car, cdr) +let match_sigma sigma ex = + match EConstr.kind sigma ex with + | App (f, [| a; p; car; cdr |]) when Termops.is_global sigma (Lazy.force coq_exist_ref) f -> + build_sigma (), (snd (destConstruct sigma f), a, p, car, cdr) + | App (f, [| a; p; car; cdr |]) when Termops.is_global sigma (Lazy.force coq_existT_ref) f -> + build_sigma_type (), (snd (destConstruct sigma f), a, p, car, cdr) | _ -> raise PatternMatchingFailure let find_sigma_data_decompose ex = (* fails with PatternMatchingFailure *) @@ -493,12 +499,12 @@ let find_sigma_data_decompose ex = (* fails with PatternMatchingFailure *) let coq_sig_pattern = lazy (mkPattern (mkGAppRef coq_sig_ref [mkGPatVar "X1"; mkGPatVar "X2"])) -let match_sigma t = - match Id.Map.bindings (matches (Lazy.force coq_sig_pattern) t) with - | [(_,a); (_,p)] -> (a,p) +let match_sigma sigma t = + match Id.Map.bindings (matches sigma (Lazy.force coq_sig_pattern) t) with + | [(_,a); (_,p)] -> (EConstr.of_constr a,EConstr.of_constr p) | _ -> anomaly (Pp.str "Unexpected pattern") -let is_matching_sigma t = is_matching (Lazy.force coq_sig_pattern) t +let is_matching_sigma sigma t = is_matching sigma (Lazy.force coq_sig_pattern) t (*** Decidable equalities *) @@ -530,15 +536,15 @@ let coq_eqdec_rev_pattern = coq_eqdec ~sum:coq_or_ref ~rev:true let op_or = coq_or_ref let op_sum = coq_sumbool_ref -let match_eqdec t = +let match_eqdec sigma t = let eqonleft,op,subst = - try true,op_sum,matches (Lazy.force coq_eqdec_inf_pattern) t + try true,op_sum,matches sigma (Lazy.force coq_eqdec_inf_pattern) t with PatternMatchingFailure -> - try false,op_sum,matches (Lazy.force coq_eqdec_inf_rev_pattern) t + try false,op_sum,matches sigma (Lazy.force coq_eqdec_inf_rev_pattern) t with PatternMatchingFailure -> - try true,op_or,matches (Lazy.force coq_eqdec_pattern) t + try true,op_or,matches sigma (Lazy.force coq_eqdec_pattern) t with PatternMatchingFailure -> - false,op_or,matches (Lazy.force coq_eqdec_rev_pattern) t in + false,op_or,matches sigma (Lazy.force coq_eqdec_rev_pattern) t in match Id.Map.bindings subst with | [(_,typ);(_,c1);(_,c2)] -> eqonleft, Universes.constr_of_global (Lazy.force op), c1, c2, typ @@ -548,8 +554,8 @@ let match_eqdec t = let coq_not_pattern = lazy (mkPattern (mkGAppRef coq_not_ref [mkGHole])) let coq_imp_False_pattern = lazy (mkPattern (mkGArrow mkGHole (mkGRef coq_False_ref))) -let is_matching_not t = is_matching (Lazy.force coq_not_pattern) t -let is_matching_imp_False t = is_matching (Lazy.force coq_imp_False_pattern) t +let is_matching_not sigma t = is_matching sigma (Lazy.force coq_not_pattern) t +let is_matching_imp_False sigma t = is_matching sigma (Lazy.force coq_imp_False_pattern) t (* Remark: patterns that have references to the standard library must be evaluated lazily (i.e. at the time they are used, not a the time diff --git a/tactics/hipattern.mli b/tactics/hipattern.mli index 8a453bf31..094d62df6 100644 --- a/tactics/hipattern.mli +++ b/tactics/hipattern.mli @@ -8,6 +8,8 @@ open Names open Term +open Evd +open EConstr open Coqlib (** High-order patterns *) @@ -40,8 +42,8 @@ open Coqlib also work on ad-hoc disjunctions introduced by the user. (Eduardo, 6/8/97). *) -type 'a matching_function = Evd.evar_map -> constr -> 'a option -type testing_function = Evd.evar_map -> constr -> bool +type 'a matching_function = evar_map -> constr -> 'a option +type testing_function = evar_map -> constr -> bool val match_with_non_recursive_type : (constr * constr list) matching_function val is_non_recursive_type : testing_function @@ -113,7 +115,7 @@ type equation_kind = exception NoEquationFound val match_with_equation: - constr -> coq_eq_data option * constr * equation_kind + evar_map -> constr -> coq_eq_data option * constr * equation_kind (***** Destructing patterns bound to some theory *) @@ -127,25 +129,25 @@ val find_this_eq_data_decompose : ([ `NF ], 'r) Proofview.Goal.t -> constr -> coq_eq_data * Univ.universe_instance * (types * constr * constr) (** A variant that returns more informative structure on the equality found *) -val find_eq_data : constr -> coq_eq_data * Univ.universe_instance * equation_kind +val find_eq_data : evar_map -> constr -> coq_eq_data * Univ.universe_instance * equation_kind (** Match a term of the form [(existT A P t p)] Returns associated lemmas and [A,P,t,p] *) -val find_sigma_data_decompose : constr -> +val find_sigma_data_decompose : evar_map -> constr -> coq_sigma_data * (Univ.universe_instance * constr * constr * constr * constr) (** Match a term of the form [{x:A|P}], returns [A] and [P] *) -val match_sigma : constr -> constr * constr +val match_sigma : evar_map -> constr -> constr * constr -val is_matching_sigma : constr -> bool +val is_matching_sigma : evar_map -> constr -> bool (** Match a decidable equality judgement (e.g [{t=u:>T}+{~t=u}]), returns [t,u,T] and a boolean telling if equality is on the left side *) -val match_eqdec : constr -> bool * constr * constr * constr * constr +val match_eqdec : evar_map -> constr -> bool * Constr.constr * Constr.constr * Constr.constr * Constr.constr (** Match an equality up to conversion; returns [(eq,t1,t2)] in normal form *) -val dest_nf_eq : ([ `NF ], 'r) Proofview.Goal.t -> constr -> (constr * constr * constr) +val dest_nf_eq : ([ `NF ], 'r) Proofview.Goal.t -> constr -> (Constr.constr * Constr.constr * Constr.constr) (** Match a negation *) -val is_matching_not : constr -> bool -val is_matching_imp_False : constr -> bool +val is_matching_not : evar_map -> constr -> bool +val is_matching_imp_False : evar_map -> constr -> bool diff --git a/tactics/inv.ml b/tactics/inv.ml index 60f1c3542..a971b9356 100644 --- a/tactics/inv.ml +++ b/tactics/inv.ml @@ -341,7 +341,7 @@ let projectAndApply as_mode thin avoid id eqname names depids = Proofview.Goal.nf_enter { enter = begin fun gl -> (** We only look at the type of hypothesis "id" *) let hyp = pf_nf_evar gl (pf_get_hyp_typ id (Proofview.Goal.assume gl)) in - let (t,t1,t2) = Hipattern.dest_nf_eq gl hyp in + let (t,t1,t2) = Hipattern.dest_nf_eq gl (EConstr.of_constr hyp) in match (kind_of_term t1, kind_of_term t2) with | Var id1, _ -> generalizeRewriteIntros as_mode (subst_hyp true id) depids id1 | _, Var id2 -> generalizeRewriteIntros as_mode (subst_hyp false id) depids id2 diff --git a/tactics/tactics.ml b/tactics/tactics.ml index f262aefa7..a04fb7ca2 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -1654,11 +1654,12 @@ let descend_in_conjunctions avoid tac (err, info) c = let t = EConstr.of_constr t in let ((ind,u),t) = reduce_to_quantified_ind env sigma t in let sign,ccl = decompose_prod_assum t in + let ccl = EConstr.of_constr ccl in match match_with_tuple sigma ccl with | Some (_,_,isrec) -> let n = (constructors_nrealargs ind).(0) in let sort = Tacticals.New.elimination_sort_of_goal gl in - let IndType (indf,_) = find_rectype env sigma (EConstr.of_constr ccl) in + let IndType (indf,_) = find_rectype env sigma ccl in let (_,inst), params = dest_ind_family indf in let cstr = (get_constructors env indf).(0) in let elim = @@ -2324,7 +2325,7 @@ let intro_decomp_eq loc l thin tac id = let t = Tacmach.New.pf_unsafe_type_of gl (EConstr.of_constr c) in let t = EConstr.of_constr t in let _,t = Tacmach.New.pf_reduce_to_quantified_ind gl t in - match my_find_eq_data_decompose gl t with + match my_find_eq_data_decompose gl (EConstr.of_constr t) with | Some (eq,u,eq_args) -> !intro_decomp_eq_function (fun n -> tac ((dloc,id)::thin) (Some (true,n)) l) @@ -2363,8 +2364,11 @@ let rewrite_hyp_then assert_style with_evars thin l2r id tac = let type_of = Tacmach.New.pf_unsafe_type_of gl in let whd_all = Tacmach.New.pf_apply whd_all gl in let t = whd_all (EConstr.of_constr (type_of (EConstr.mkVar id))) in + let t = EConstr.of_constr t in let eqtac, thin = match match_with_equality_type sigma t with | Some (hdcncl,[_;lhs;rhs]) -> + let lhs = EConstr.Unsafe.to_constr lhs in + let rhs = EConstr.Unsafe.to_constr rhs in if l2r && isVar lhs && not (occur_var env sigma (destVar lhs) (EConstr.of_constr rhs)) then let id' = destVar lhs in subst_on l2r id' rhs, early_clear id' thin @@ -2375,6 +2379,7 @@ let rewrite_hyp_then assert_style with_evars thin l2r id tac = Tacticals.New.tclTHEN (rew_on l2r onConcl) (clear [id]), thin | Some (hdcncl,[c]) -> + let c = EConstr.Unsafe.to_constr c in let l2r = not l2r in (* equality of the form eq_true *) if isVar c then let id' = destVar c in @@ -4689,6 +4694,7 @@ let reflexivity_red allowred = inside setoid_reflexivity (see Optimize cases in setoid_replace.ml). *) let sigma = Tacmach.New.project gl in let concl = maybe_betadeltaiota_concl allowred gl in + let concl = EConstr.of_constr concl in match match_with_equality_type sigma concl with | None -> Proofview.tclZERO NoEquationFound | Some _ -> one_constructor 1 NoBindings @@ -4716,19 +4722,21 @@ let (forward_setoid_symmetry, setoid_symmetry) = Hook.make () (* This is probably not very useful any longer *) let prove_symmetry hdcncl eq_kind = let symc = + let open EConstr in match eq_kind with | MonomorphicLeibnizEq (c1,c2) -> mkApp(hdcncl,[|c2;c1|]) | PolymorphicLeibnizEq (typ,c1,c2) -> mkApp(hdcncl,[|typ;c2;c1|]) | HeterogenousEq (t1,c1,t2,c2) -> mkApp(hdcncl,[|t2;c2;t1;c1|]) in + let symc = EConstr.Unsafe.to_constr symc in Tacticals.New.tclTHENFIRST (cut symc) (Tacticals.New.tclTHENLIST [ intro; Tacticals.New.onLastHyp simplest_case; one_constructor 1 NoBindings ]) -let match_with_equation c = +let match_with_equation sigma c = try - let res = match_with_equation c in + let res = match_with_equation sigma c in Proofview.tclUNIT res with NoEquationFound -> Proofview.tclZERO NoEquationFound @@ -4738,8 +4746,9 @@ let symmetry_red allowred = (* PL: usual symmetry don't perform any reduction when searching for an equality, but we may need to do some when called back from inside setoid_reflexivity (see Optimize cases in setoid_replace.ml). *) + let sigma = Tacmach.New.project gl in let concl = maybe_betadeltaiota_concl allowred gl in - match_with_equation concl >>= fun with_eqn -> + match_with_equation sigma (EConstr.of_constr concl) >>= fun with_eqn -> match with_eqn with | Some eq_data,_,_ -> Tacticals.New.tclTHEN @@ -4761,15 +4770,20 @@ let (forward_setoid_symmetry_in, setoid_symmetry_in) = Hook.make () let symmetry_in id = Proofview.Goal.enter { enter = begin fun gl -> + let sigma = Tacmach.New.project gl in let ctype = Tacmach.New.pf_unsafe_type_of gl (EConstr.mkVar id) in let sign,t = decompose_prod_assum ctype in + let t = EConstr.of_constr t in Proofview.tclORELSE begin - match_with_equation t >>= fun (_,hdcncl,eq) -> - let symccl = match eq with + match_with_equation sigma t >>= fun (_,hdcncl,eq) -> + let symccl = + let open EConstr in + match eq with | MonomorphicLeibnizEq (c1,c2) -> mkApp (hdcncl, [| c2; c1 |]) | PolymorphicLeibnizEq (typ,c1,c2) -> mkApp (hdcncl, [| typ; c2; c1 |]) | HeterogenousEq (t1,c1,t2,c2) -> mkApp (hdcncl, [| t2; c2; t1; c1 |]) in + let symccl = EConstr.Unsafe.to_constr symccl in Tacticals.New.tclTHENS (cut (it_mkProd_or_LetIn symccl sign)) [ intro_replacing id; Tacticals.New.tclTHENLIST [ intros; symmetry; apply (mkVar id); assumption ] ] @@ -4804,6 +4818,8 @@ let (forward_setoid_transitivity, setoid_transitivity) = Hook.make () (* This is probably not very useful any longer *) let prove_transitivity hdcncl eq_kind t = Proofview.Goal.enter { enter = begin fun gl -> + let t = EConstr.of_constr t in + let open EConstr in let (eq1,eq2) = match eq_kind with | MonomorphicLeibnizEq (c1,c2) -> mkApp (hdcncl, [| c1; t|]), mkApp (hdcncl, [| t; c2 |]) @@ -4813,10 +4829,13 @@ let prove_transitivity hdcncl eq_kind t = let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in let type_of = Typing.unsafe_type_of env sigma in - let typt = type_of (EConstr.of_constr t) in + let typt = type_of t in + let typt = EConstr.of_constr typt in (mkApp(hdcncl, [| typ1; c1; typt ;t |]), mkApp(hdcncl, [| typt; t; typ2; c2 |])) in + let eq1 = EConstr.Unsafe.to_constr eq1 in + let eq2 = EConstr.Unsafe.to_constr eq2 in Tacticals.New.tclTHENFIRST (cut eq2) (Tacticals.New.tclTHENFIRST (cut eq1) (Tacticals.New.tclTHENLIST @@ -4830,8 +4849,9 @@ let transitivity_red allowred t = (* PL: usual transitivity don't perform any reduction when searching for an equality, but we may need to do some when called back from inside setoid_reflexivity (see Optimize cases in setoid_replace.ml). *) + let sigma = Tacmach.New.project gl in let concl = maybe_betadeltaiota_concl allowred gl in - match_with_equation concl >>= fun with_eqn -> + match_with_equation sigma (EConstr.of_constr concl) >>= fun with_eqn -> match with_eqn with | Some eq_data,_,_ -> Tacticals.New.tclTHEN diff --git a/tactics/tactics.mli b/tactics/tactics.mli index 268453152..368a1df76 100644 --- a/tactics/tactics.mli +++ b/tactics/tactics.mli @@ -413,7 +413,7 @@ val subst_one : val declare_intro_decomp_eq : ((int -> unit Proofview.tactic) -> Coqlib.coq_eq_data * types * - (types * constr * constr) -> + (EConstr.types * EConstr.constr * EConstr.constr) -> constr * types -> unit Proofview.tactic) -> unit (** {6 Simple form of basic tactics. } *) diff --git a/toplevel/indschemes.ml b/toplevel/indschemes.ml index 0723f16c3..9ed34d713 100644 --- a/toplevel/indschemes.ml +++ b/toplevel/indschemes.ml @@ -313,7 +313,7 @@ let warn_cannot_build_congruence = strbrk "Cannot build congruence scheme because eq is not found") let declare_congr_scheme ind = - if Hipattern.is_equality_type Evd.empty (mkInd ind) (** FIXME *) then begin + if Hipattern.is_equality_type Evd.empty (EConstr.of_constr (mkInd ind)) (** FIXME *) then begin if try Coqlib.check_required_library Coqlib.logic_module_name; true with e when CErrors.noncritical e -> false -- cgit v1.2.3 From 485bbfbed4ae4a28119c4e42c5e40fd77abf4f8a Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sun, 13 Nov 2016 20:38:41 +0100 Subject: Tactics API using EConstr. --- engine/eConstr.ml | 22 + engine/eConstr.mli | 8 + engine/evarutil.ml | 2 + engine/evarutil.mli | 12 +- engine/proofview.ml | 1 + engine/termops.ml | 22 +- engine/termops.mli | 12 +- interp/stdarg.mli | 4 +- intf/tactypes.mli | 4 +- ltac/coretactics.ml4 | 20 +- ltac/extratactics.ml4 | 23 +- ltac/g_auto.ml4 | 12 +- ltac/g_rewrite.ml4 | 2 +- ltac/pptactic.ml | 17 +- ltac/rewrite.ml | 35 +- ltac/rewrite.mli | 2 +- ltac/tacexpr.mli | 6 +- ltac/tacinterp.ml | 29 +- ltac/tauto.ml | 10 +- plugins/btauto/refl_btauto.ml | 3 +- plugins/cc/cctac.ml | 20 +- plugins/decl_mode/decl_proof_instr.ml | 37 +- plugins/decl_mode/decl_proof_instr.mli | 2 +- plugins/firstorder/instances.ml | 18 +- plugins/firstorder/rules.ml | 27 +- plugins/fourier/fourierR.ml | 31 +- plugins/funind/functional_principles_proofs.ml | 42 +- plugins/funind/functional_principles_types.ml | 10 +- plugins/funind/g_indfun.ml4 | 10 +- plugins/funind/indfun.ml | 7 +- plugins/funind/invfun.ml | 44 +- plugins/funind/merge.ml | 9 +- plugins/funind/recdef.ml | 83 +-- plugins/micromega/coq_micromega.ml | 14 +- plugins/nsatz/nsatz.ml | 1 + plugins/omega/coq_omega.ml | 39 +- plugins/quote/quote.ml | 4 +- plugins/romega/refl_omega.ml | 9 +- plugins/rtauto/refl_tauto.ml | 1 + plugins/ssrmatching/ssrmatching.ml4 | 1 + pretyping/cases.ml | 9 +- pretyping/classops.ml | 1 + pretyping/coercion.ml | 2 +- pretyping/detyping.ml | 2 +- pretyping/evarconv.ml | 8 +- pretyping/evardefine.ml | 17 +- pretyping/evarsolve.ml | 13 +- pretyping/inductiveops.ml | 1 + pretyping/inductiveops.mli | 2 +- pretyping/pretyping.ml | 6 +- pretyping/tacred.ml | 8 +- pretyping/tacred.mli | 10 +- pretyping/unification.ml | 9 +- proofs/clenv.ml | 2 - proofs/tacmach.mli | 6 +- stm/lemmas.ml | 4 +- stm/stm.ml | 2 +- tactics/auto.ml | 1 + tactics/class_tactics.ml | 6 +- tactics/contradiction.ml | 14 +- tactics/contradiction.mli | 2 +- tactics/eauto.ml | 10 +- tactics/elim.ml | 4 +- tactics/eqdecide.ml | 19 +- tactics/eqschemes.ml | 2 +- tactics/equality.ml | 63 +- tactics/equality.mli | 16 +- tactics/hints.ml | 2 +- tactics/inv.ml | 8 +- tactics/inv.mli | 2 +- tactics/leminv.ml | 4 +- tactics/leminv.mli | 2 +- tactics/tacticals.ml | 5 +- tactics/tacticals.mli | 8 +- tactics/tactics.ml | 832 +++++++++++++------------ tactics/tactics.mli | 7 +- tactics/term_dnet.ml | 2 +- toplevel/auto_ind_decl.ml | 40 +- toplevel/command.ml | 4 +- 79 files changed, 998 insertions(+), 812 deletions(-) diff --git a/engine/eConstr.ml b/engine/eConstr.ml index 7bd708e31..9e0a55a0d 100644 --- a/engine/eConstr.ml +++ b/engine/eConstr.ml @@ -98,6 +98,7 @@ 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) @@ -466,6 +467,11 @@ let eq_constr_nounivs sigma 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 @@ -608,6 +614,22 @@ let mkLambda_or_LetIn decl c = | LocalAssum (na,t) -> mkLambda (na, of_constr t, c) | LocalDef (na,b,t) -> mkLetIn (na, of_constr b, of_constr 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 (of_constr t) c + | LocalDef (id,b,t) -> mkNamedLetIn id (of_constr b) (of_constr t) c + +let mkNamedLambda_or_LetIn decl c = + let open Context.Named.Declaration in + match decl with + | LocalAssum (id,t) -> mkNamedLambda id (of_constr t) c + | LocalDef (id,b,t) -> mkNamedLetIn id (of_constr b) (of_constr 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 diff --git a/engine/eConstr.mli b/engine/eConstr.mli index e4136a612..15463a8f6 100644 --- a/engine/eConstr.mli +++ b/engine/eConstr.mli @@ -73,6 +73,7 @@ val mkConstructU : pconstructor -> 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 @@ -81,6 +82,12 @@ val mkLambda_or_LetIn : Rel.Declaration.t -> t -> t val it_mkProd_or_LetIn : t -> Rel.t -> t val it_mkLambda_or_LetIn : t -> Rel.t -> 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.t -> types -> types +val mkNamedProd_or_LetIn : Named.Declaration.t -> types -> types + (** {6 Simple case analysis} *) val isRel : Evd.evar_map -> t -> bool @@ -141,6 +148,7 @@ 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} *) diff --git a/engine/evarutil.ml b/engine/evarutil.ml index 7ccf9d810..4f40499d0 100644 --- a/engine/evarutil.ml +++ b/engine/evarutil.ml @@ -367,6 +367,7 @@ let push_rel_decl_to_named_context decl (subst, vsubst, avoid, nc) = let push_rel_context_to_named_context env 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 @@ -421,6 +422,7 @@ let new_pure_evar sign evd ?(src=default_source) ?(filter = Filter.identity) ?ca Sigma.Unsafe.of_pair (newevk, evd) 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 diff --git a/engine/evarutil.mli b/engine/evarutil.mli index 431d98083..6620bbaed 100644 --- a/engine/evarutil.mli +++ b/engine/evarutil.mli @@ -24,7 +24,7 @@ val new_evar : env -> 'r Sigma.t -> ?src:Loc.t * Evar_kinds.t -> ?filter:Filter.t -> ?candidates:constr list -> ?store:Store.t -> ?naming:Misctypes.intro_pattern_naming_expr -> - ?principal:bool -> EConstr.types -> (constr, 'r) Sigma.sigma + ?principal:bool -> EConstr.types -> (EConstr.constr, 'r) Sigma.sigma val new_pure_evar : named_context_val -> 'r Sigma.t -> ?src:Loc.t * Evar_kinds.t -> ?filter:Filter.t -> @@ -39,18 +39,18 @@ val e_new_evar : env -> evar_map ref -> ?src:Loc.t * Evar_kinds.t -> ?filter:Filter.t -> ?candidates:constr list -> ?store:Store.t -> ?naming:Misctypes.intro_pattern_naming_expr -> - ?principal:bool -> EConstr.types -> constr + ?principal:bool -> EConstr.types -> EConstr.constr (** 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 -> ?naming:Misctypes.intro_pattern_naming_expr -> ?principal:bool -> rigid -> - (constr * sorts, 'r) Sigma.sigma + (EConstr.constr * sorts, 'r) Sigma.sigma val e_new_type_evar : env -> evar_map ref -> ?src:Loc.t * Evar_kinds.t -> ?filter:Filter.t -> - ?naming:Misctypes.intro_pattern_naming_expr -> ?principal:bool -> rigid -> constr * sorts + ?naming:Misctypes.intro_pattern_naming_expr -> ?principal:bool -> rigid -> EConstr.constr * sorts val new_Type : ?rigid:rigid -> env -> 'r Sigma.t -> (constr, 'r) Sigma.sigma val e_new_Type : ?rigid:rigid -> env -> evar_map ref -> constr @@ -74,7 +74,7 @@ val new_evar_instance : ?src:Loc.t * Evar_kinds.t -> ?filter:Filter.t -> ?candidates:constr list -> ?store:Store.t -> ?naming:Misctypes.intro_pattern_naming_expr -> ?principal:bool -> - constr list -> (constr, 'r) Sigma.sigma + EConstr.constr list -> (EConstr.constr, 'r) Sigma.sigma val make_pure_subst : evar_info -> 'a array -> (Id.t * 'a) list @@ -218,7 +218,7 @@ val push_rel_decl_to_named_context : Context.Rel.Declaration.t -> ext_named_context -> ext_named_context val push_rel_context_to_named_context : Environ.env -> EConstr.types -> - named_context_val * EConstr.types * constr list * csubst * (identifier*EConstr.constr) list + named_context_val * EConstr.types * EConstr.constr list * csubst * (identifier*EConstr.constr) list val generalize_evar_over_rels : evar_map -> existential -> types * constr list diff --git a/engine/proofview.ml b/engine/proofview.ml index b0f6d463b..9adf94744 100644 --- a/engine/proofview.ml +++ b/engine/proofview.ml @@ -72,6 +72,7 @@ let dependent_init = | 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 (EConstr.of_constr typ) in + let econstr = EConstr.Unsafe.to_constr econstr in let sigma = Sigma.to_evar_map sigma in let ret, { solution = sol; comb = comb } = aux (t sigma econstr) in let (gl, _) = Term.destEvar econstr in diff --git a/engine/termops.ml b/engine/termops.ml index 59dbb73f5..b7932665a 100644 --- a/engine/termops.ml +++ b/engine/termops.ml @@ -159,6 +159,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 @@ -857,16 +858,18 @@ let base_sort_cmp pb s0 s1 = | _ -> 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 +let compare_constr_univ sigma f cv_pb t1 t2 = + match EConstr.kind sigma t1, EConstr.kind sigma t2 with Sort s1, Sort s2 -> base_sort_cmp cv_pb s1 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 + | _ -> 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) *) @@ -883,12 +886,12 @@ type subst = (Context.Rel.t * constr) Evar.Map.t exception CannotFilter -let filtering env cv_pb c1 c2 = +let filtering sigma env cv_pb c1 c2 = 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 (EConstr.of_constr c1) (EConstr.of_constr (lift shift c2)) then () else raise CannotFilter with Not_found -> evm := Evar.Map.add ev (e1,c1) !evm in @@ -909,8 +912,9 @@ 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 - (fun pb c1 c2 -> aux env pb c1 c2; true) cv_pb c1 c2 then () + let inj = EConstr.Unsafe.to_constr in + if compare_constr_univ sigma + (fun pb c1 c2 -> aux env pb (inj c1) (inj c2); true) cv_pb (EConstr.of_constr c1) (EConstr.of_constr c2) then () else raise CannotFilter (* TODO: le reste des binders *) in diff --git a/engine/termops.mli b/engine/termops.mli index abc9caa98..7758a57ee 100644 --- a/engine/termops.mli +++ b/engine/termops.mli @@ -41,7 +41,7 @@ val lookup_rel_id : Id.t -> Context.Rel.t -> int * constr option * types [rel_vect n m] builds [|Rel (n+m);...;Rel(n+1)|] *) val rel_vect : int -> int -> constr array -val rel_list : int -> int -> constr list +val rel_list : int -> int -> EConstr.constr list (** iterators/destructors on terms *) val mkProd_or_LetIn : Context.Rel.Declaration.t -> types -> types @@ -160,10 +160,10 @@ val replace_term : Evd.evar_map -> EConstr.t -> EConstr.t -> EConstr.t -> constr (** Alternative term equalities *) val base_sort_cmp : Reduction.conv_pb -> sorts -> sorts -> bool -val compare_constr_univ : (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 compare_constr_univ : Evd.evar_map -> (Reduction.conv_pb -> EConstr.constr -> EConstr.constr -> bool) -> + Reduction.conv_pb -> EConstr.constr -> EConstr.constr -> bool +val constr_cmp : Evd.evar_map -> Reduction.conv_pb -> EConstr.constr -> EConstr.constr -> bool +val eq_constr : Evd.evar_map -> EConstr.constr -> EConstr.constr -> bool (* FIXME rename: erases universes*) val eta_reduce_head : constr -> constr @@ -185,7 +185,7 @@ exception 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 +val filtering : Evd.evar_map -> Context.Rel.t -> Reduction.conv_pb -> constr -> constr -> subst val decompose_prod_letin : Evd.evar_map -> EConstr.t -> int * Context.Rel.t * constr val align_prod_letin : Evd.evar_map -> EConstr.t -> EConstr.t -> Context.Rel.t * constr diff --git a/interp/stdarg.mli b/interp/stdarg.mli index af3a73462..3047d2bce 100644 --- a/interp/stdarg.mli +++ b/interp/stdarg.mli @@ -57,12 +57,12 @@ val wit_open_constr : val wit_constr_with_bindings : (constr_expr with_bindings, glob_constr_and_expr with_bindings, - constr with_bindings delayed_open) genarg_type + EConstr.constr with_bindings delayed_open) genarg_type val wit_bindings : (constr_expr bindings, glob_constr_and_expr bindings, - constr bindings delayed_open) genarg_type + EConstr.constr bindings delayed_open) genarg_type val wit_red_expr : ((constr_expr,reference or_by_notation,constr_expr) red_expr_gen, diff --git a/intf/tactypes.mli b/intf/tactypes.mli index b96cb67df..02cfc44e2 100644 --- a/intf/tactypes.mli +++ b/intf/tactypes.mli @@ -26,8 +26,8 @@ type glob_constr_pattern_and_expr = Id.Set.t * glob_constr_and_expr * constr_pat type 'a delayed_open = { delayed : 'r. Environ.env -> 'r Sigma.t -> ('a, 'r) Sigma.sigma } -type delayed_open_constr = Term.constr delayed_open -type delayed_open_constr_with_bindings = Term.constr with_bindings delayed_open +type delayed_open_constr = EConstr.constr delayed_open +type delayed_open_constr_with_bindings = EConstr.constr with_bindings delayed_open type intro_pattern = delayed_open_constr intro_pattern_expr located type intro_patterns = delayed_open_constr intro_pattern_expr located list diff --git a/ltac/coretactics.ml4 b/ltac/coretactics.ml4 index 28ff6df83..20d9640fc 100644 --- a/ltac/coretactics.ml4 +++ b/ltac/coretactics.ml4 @@ -27,7 +27,7 @@ TACTIC EXTEND reflexivity END TACTIC EXTEND exact - [ "exact" casted_constr(c) ] -> [ Tactics.exact_no_check c ] + [ "exact" casted_constr(c) ] -> [ Tactics.exact_no_check (EConstr.of_constr c) ] END TACTIC EXTEND assumption @@ -39,35 +39,35 @@ TACTIC EXTEND etransitivity END TACTIC EXTEND cut - [ "cut" constr(c) ] -> [ Tactics.cut c ] + [ "cut" constr(c) ] -> [ Tactics.cut (EConstr.of_constr c) ] END TACTIC EXTEND exact_no_check - [ "exact_no_check" constr(c) ] -> [ Tactics.exact_no_check c ] + [ "exact_no_check" constr(c) ] -> [ Tactics.exact_no_check (EConstr.of_constr c) ] END TACTIC EXTEND vm_cast_no_check - [ "vm_cast_no_check" constr(c) ] -> [ Tactics.vm_cast_no_check c ] + [ "vm_cast_no_check" constr(c) ] -> [ Tactics.vm_cast_no_check (EConstr.of_constr c) ] END TACTIC EXTEND native_cast_no_check - [ "native_cast_no_check" constr(c) ] -> [ Tactics.native_cast_no_check c ] + [ "native_cast_no_check" constr(c) ] -> [ Tactics.native_cast_no_check (EConstr.of_constr c) ] END TACTIC EXTEND casetype - [ "casetype" constr(c) ] -> [ Tactics.case_type c ] + [ "casetype" constr(c) ] -> [ Tactics.case_type (EConstr.of_constr c) ] END TACTIC EXTEND elimtype - [ "elimtype" constr(c) ] -> [ Tactics.elim_type c ] + [ "elimtype" constr(c) ] -> [ Tactics.elim_type (EConstr.of_constr c) ] END TACTIC EXTEND lapply - [ "lapply" constr(c) ] -> [ Tactics.cut_and_apply c ] + [ "lapply" constr(c) ] -> [ Tactics.cut_and_apply (EConstr.of_constr c) ] END TACTIC EXTEND transitivity - [ "transitivity" constr(c) ] -> [ Tactics.intros_transitivity (Some c) ] + [ "transitivity" constr(c) ] -> [ Tactics.intros_transitivity (Some (EConstr.of_constr c)) ] END (** Left *) @@ -297,7 +297,7 @@ END (* Generalize dependent *) TACTIC EXTEND generalize_dependent - [ "generalize" "dependent" constr(c) ] -> [ Tactics.generalize_dep c ] + [ "generalize" "dependent" constr(c) ] -> [ Tactics.generalize_dep (EConstr.of_constr c) ] END (* Table of "pervasives" macros tactics (e.g. auto, simpl, etc.) *) diff --git a/ltac/extratactics.ml4 b/ltac/extratactics.ml4 index 3e7cf5d13..c39b1a0e9 100644 --- a/ltac/extratactics.ml4 +++ b/ltac/extratactics.ml4 @@ -118,7 +118,7 @@ END let discrHyp id = Proofview.tclEVARMAP >>= fun sigma -> - discr_main { delayed = fun env sigma -> Sigma.here (Term.mkVar id, NoBindings) sigma } + discr_main { delayed = fun env sigma -> Sigma.here (EConstr.mkVar id, NoBindings) sigma } let injection_main with_evars c = elimOnConstrWithHoles (injClause None) with_evars c @@ -150,7 +150,7 @@ END let injHyp id = Proofview.tclEVARMAP >>= fun sigma -> - injection_main false { delayed = fun env sigma -> Sigma.here (Term.mkVar id, NoBindings) sigma } + injection_main false { delayed = fun env sigma -> Sigma.here (EConstr.mkVar id, NoBindings) sigma } TACTIC EXTEND dependent_rewrite | [ "dependent" "rewrite" orient(b) constr(c) ] -> [ rewriteInConcl b c ] @@ -301,6 +301,7 @@ let project_hint pri l2r r = let t = Retyping.get_type_of env sigma (EConstr.of_constr c) in let t = Tacred.reduce_to_quantified_ref env sigma (Lazy.force coq_iff_ref) (EConstr.of_constr t) in + let t = EConstr.Unsafe.to_constr t in let sign,ccl = decompose_prod_assum t in let (a,b) = match snd (decompose_app ccl) with | [a;b] -> (a,b) @@ -475,6 +476,7 @@ let transitivity_left_table = Summary.ref [] ~name:"transitivity-steps-l" let step left x tac = let l = List.map (fun lem -> + let lem = EConstr.of_constr lem in Tacticals.New.tclTHENLAST (apply_with_bindings (lem, ImplicitBindings [x])) tac) @@ -510,13 +512,13 @@ let add_transitivity_lemma left lem = (* Vernacular syntax *) TACTIC EXTEND stepl -| ["stepl" constr(c) "by" tactic(tac) ] -> [ step true c (Tacinterp.tactic_of_value ist tac) ] -| ["stepl" constr(c) ] -> [ step true c (Proofview.tclUNIT ()) ] +| ["stepl" constr(c) "by" tactic(tac) ] -> [ step true (EConstr.of_constr c) (Tacinterp.tactic_of_value ist tac) ] +| ["stepl" constr(c) ] -> [ step true (EConstr.of_constr c) (Proofview.tclUNIT ()) ] END TACTIC EXTEND stepr -| ["stepr" constr(c) "by" tactic(tac) ] -> [ step false c (Tacinterp.tactic_of_value ist tac) ] -| ["stepr" constr(c) ] -> [ step false c (Proofview.tclUNIT ()) ] +| ["stepr" constr(c) "by" tactic(tac) ] -> [ step false (EConstr.of_constr c) (Tacinterp.tactic_of_value ist tac) ] +| ["stepr" constr(c) ] -> [ step false (EConstr.of_constr c) (Proofview.tclUNIT ()) ] END VERNAC COMMAND EXTEND AddStepl CLASSIFIED AS SIDEFF @@ -660,7 +662,7 @@ let hResolve id c occ t = let sigma = Evd.merge_universe_context sigma ctx in let t_constr_type = Retyping.get_type_of env sigma (EConstr.of_constr t_constr) in let tac = - (change_concl (mkLetIn (Anonymous,t_constr,t_constr_type,concl))) + (change_concl (EConstr.of_constr (mkLetIn (Anonymous,t_constr,t_constr_type,concl)))) in Sigma.Unsafe.of_pair (tac, sigma) end } @@ -694,7 +696,7 @@ let hget_evar n = if n <= 0 then error "Incorrect existential variable index."; let ev = List.nth evl (n-1) in let ev_type = existential_type sigma ev in - change_concl (mkLetIn (Anonymous,mkEvar ev,ev_type,concl)) + change_concl (EConstr.of_constr (mkLetIn (Anonymous,mkEvar ev,ev_type,concl))) end } TACTIC EXTEND hget_evar @@ -736,15 +738,16 @@ let mkCaseEq a : unit Proofview.tactic = Proofview.Goal.nf_enter { enter = begin fun gl -> let type_of_a = Tacmach.New.of_old (fun g -> Tacmach.pf_unsafe_type_of g (EConstr.of_constr a)) gl in Tacticals.New.tclTHENLIST - [Tactics.generalize [mkApp(delayed_force refl_equal, [| type_of_a; a|])]; + [Tactics.generalize [EConstr.of_constr (mkApp(delayed_force refl_equal, [| type_of_a; a|]))]; Proofview.Goal.nf_enter { enter = begin fun gl -> let concl = Proofview.Goal.concl gl in let env = Proofview.Goal.env gl in (** FIXME: this looks really wrong. Does anybody really use this tactic? *) let Sigma (c, _, _) = (Tacred.pattern_occs [Locus.OnlyOccurrences [1], EConstr.of_constr a]).Reductionops.e_redfun env (Sigma.Unsafe.of_evar_map Evd.empty) (EConstr.of_constr concl) in + let c = EConstr.of_constr c in change_concl c end }; - simplest_case a] + simplest_case (EConstr.of_constr a)] end } diff --git a/ltac/g_auto.ml4 b/ltac/g_auto.ml4 index 82ba63871..c6395d7e2 100644 --- a/ltac/g_auto.ml4 +++ b/ltac/g_auto.ml4 @@ -48,7 +48,11 @@ let eval_uconstrs ist cs = fail_evar = false; expand_evars = true } in - List.map (fun c -> Pretyping.type_uconstr ~flags ist c) cs + let map c = { delayed = fun env sigma -> + let Sigma.Sigma (c, sigma, p) = c.delayed env sigma in + Sigma.Sigma (EConstr.of_constr c, sigma, p) + } in + List.map (fun c -> map (Pretyping.type_uconstr ~flags ist c)) cs let pr_auto_using _ _ _ = Pptactic.pr_auto_using (fun _ -> mt ()) @@ -153,7 +157,7 @@ TACTIC EXTEND autounfoldify END TACTIC EXTEND unify -| ["unify" constr(x) constr(y) ] -> [ Tactics.unify x y ] +| ["unify" constr(x) constr(y) ] -> [ Tactics.unify (EConstr.of_constr x) (EConstr.of_constr y) ] | ["unify" constr(x) constr(y) "with" preident(base) ] -> [ let table = try Some (Hints.searchtable_map base) with Not_found -> None in match table with @@ -162,13 +166,13 @@ TACTIC EXTEND unify Tacticals.New.tclZEROMSG msg | Some t -> let state = Hints.Hint_db.transparent_state t in - Tactics.unify ~state x y + Tactics.unify ~state (EConstr.of_constr x) (EConstr.of_constr y) ] END TACTIC EXTEND convert_concl_no_check -| ["convert_concl_no_check" constr(x) ] -> [ Tactics.convert_concl_no_check x Term.DEFAULTcast ] +| ["convert_concl_no_check" constr(x) ] -> [ Tactics.convert_concl_no_check (EConstr.of_constr x) Term.DEFAULTcast ] END let pr_hints_path_atom _ _ _ = Hints.pp_hints_path_atom diff --git a/ltac/g_rewrite.ml4 b/ltac/g_rewrite.ml4 index b1c4f58eb..bae5a516c 100644 --- a/ltac/g_rewrite.ml4 +++ b/ltac/g_rewrite.ml4 @@ -265,7 +265,7 @@ TACTIC EXTEND setoid_reflexivity END TACTIC EXTEND setoid_transitivity - [ "setoid_transitivity" constr(t) ] -> [ setoid_transitivity (Some t) ] + [ "setoid_transitivity" constr(t) ] -> [ setoid_transitivity (Some (EConstr.of_constr t)) ] | [ "setoid_etransitivity" ] -> [ setoid_transitivity None ] END diff --git a/ltac/pptactic.ml b/ltac/pptactic.ml index 6230fa060..934830f4d 100644 --- a/ltac/pptactic.ml +++ b/ltac/pptactic.ml @@ -1158,11 +1158,12 @@ module Make let pr_glob_tactic env = pr_glob_tactic_level env ltop let strip_prod_binders_constr n ty = + let ty = EConstr.Unsafe.to_constr ty in let rec strip_ty acc n ty = - if n=0 then (List.rev acc, ty) else + if n=0 then (List.rev acc, EConstr.of_constr ty) else match Term.kind_of_term ty with Term.Prod(na,a,b) -> - strip_ty (([Loc.ghost,na],a)::acc) (n-1) b + strip_ty (([Loc.ghost,na],EConstr.of_constr a)::acc) (n-1) b | _ -> error "Cannot translate fix tactic: not enough products" in strip_ty [] n ty @@ -1170,9 +1171,9 @@ module Make let prtac n (t:atomic_tactic_expr) = let pr = { pr_tactic = (fun _ _ -> str ""); - pr_constr = pr_constr_env env Evd.empty; + pr_constr = (fun c -> pr_constr_env env Evd.empty (EConstr.Unsafe.to_constr c)); pr_dconstr = pr_and_constr_expr (pr_glob_constr_env env); - pr_lconstr = pr_lconstr_env env Evd.empty; + pr_lconstr = (fun c -> pr_lconstr_env env Evd.empty (EConstr.Unsafe.to_constr c)); pr_pattern = pr_constr_pattern_env env Evd.empty; pr_lpattern = pr_lconstr_pattern_env env Evd.empty; pr_constant = pr_evaluable_reference_env env; @@ -1284,7 +1285,7 @@ let () = wit_intro_pattern (Miscprint.pr_intro_pattern pr_constr_expr) (Miscprint.pr_intro_pattern (fun (c,_) -> pr_glob_constr c)) - (Miscprint.pr_intro_pattern (fun c -> pr_constr (fst (run_delayed c)))); + (Miscprint.pr_intro_pattern (fun c -> pr_constr (EConstr.Unsafe.to_constr (fst (run_delayed c))))); Genprint.register_print0 wit_clause_dft_concl (pr_clauses (Some true) pr_lident) @@ -1317,15 +1318,15 @@ let () = Genprint.register_print0 wit_bindings (pr_bindings_no_with pr_constr_expr pr_lconstr_expr) (pr_bindings_no_with (pr_and_constr_expr pr_glob_constr) (pr_and_constr_expr pr_lglob_constr)) - (fun it -> pr_bindings_no_with pr_constr pr_lconstr (fst (run_delayed it))); + (fun it -> pr_bindings_no_with (EConstr.Unsafe.to_constr %> pr_constr) (EConstr.Unsafe.to_constr %> pr_lconstr) (fst (run_delayed it))); Genprint.register_print0 wit_constr_with_bindings (pr_with_bindings pr_constr_expr pr_lconstr_expr) (pr_with_bindings (pr_and_constr_expr pr_glob_constr) (pr_and_constr_expr pr_lglob_constr)) - (fun it -> pr_with_bindings pr_constr pr_lconstr (fst (run_delayed it))); + (fun it -> pr_with_bindings (EConstr.Unsafe.to_constr %> pr_constr) (EConstr.Unsafe.to_constr %> pr_lconstr) (fst (run_delayed it))); Genprint.register_print0 Tacarg.wit_destruction_arg (pr_destruction_arg pr_constr_expr pr_lconstr_expr) (pr_destruction_arg (pr_and_constr_expr pr_glob_constr) (pr_and_constr_expr pr_lglob_constr)) - (fun it -> pr_destruction_arg pr_constr pr_lconstr (run_delayed_destruction_arg it)); + (fun it -> pr_destruction_arg (EConstr.Unsafe.to_constr %> pr_constr) (EConstr.Unsafe.to_constr %> pr_lconstr) (run_delayed_destruction_arg it)); Genprint.register_print0 Stdarg.wit_int int int int; Genprint.register_print0 Stdarg.wit_bool pr_bool pr_bool pr_bool; Genprint.register_print0 Stdarg.wit_unit pr_unit pr_unit pr_unit; diff --git a/ltac/rewrite.ml b/ltac/rewrite.ml index 52cf1ff35..ef2ab0917 100644 --- a/ltac/rewrite.ml +++ b/ltac/rewrite.ml @@ -97,8 +97,8 @@ let new_cstr_evar (evd,cstrs) env t = let evd = Sigma.Unsafe.of_evar_map evd in let Sigma (t, evd', _) = Evarutil.new_evar ~store:s env evd (EConstr.of_constr t) in let evd' = Sigma.to_evar_map evd' in - let ev, _ = destEvar t in - (evd', Evar.Set.add ev cstrs), t + let ev, _ = EConstr.destEvar evd' t in + (evd', Evar.Set.add ev cstrs), EConstr.Unsafe.to_constr t (** Building or looking up instances. *) let e_new_cstr_evar env evars t = @@ -363,6 +363,7 @@ end) = struct let env' = Environ.push_rel_context rels env in let sigma = Sigma.Unsafe.of_evar_map sigma in let Sigma ((evar, _), evars, _) = Evarutil.new_type_evar env' sigma Evd.univ_flexible in + let evar = EConstr.Unsafe.to_constr evar in let evars = Sigma.to_evar_map evars in let evars, inst = app_poly env (evars,Evar.Set.empty) @@ -774,7 +775,7 @@ let poly_subrelation sort = if sort then PropGlobal.subrelation else TypeGlobal.subrelation let resolve_subrelation env avoid car rel sort prf rel' res = - if eq_constr rel rel' then res + if Termops.eq_constr (fst res.rew_evars) (EConstr.of_constr rel) (EConstr.of_constr rel') then res else let evars, app = app_poly_check env res.rew_evars (poly_subrelation sort) [|car; rel; rel'|] in let evars, subrel = new_cstr_evar evars env app in @@ -872,7 +873,7 @@ let apply_rule unify loccs : int pure_strategy = | Some rew -> let occ = succ occ in if not (is_occ occ) then (occ, Fail) - else if eq_constr t rew.rew_to then (occ, Identity) + else if Termops.eq_constr (fst rew.rew_evars) (EConstr.of_constr t) (EConstr.of_constr rew.rew_to) then (occ, Identity) else let res = { rew with rew_car = ty } in let rel, prf = get_rew_prf res in @@ -1111,7 +1112,7 @@ let subterm all flags (s : 'a pure_strategy) : 'a pure_strategy = | Prod (n, dom, codom) -> let lam = mkLambda (n, dom, codom) in let (evars', app), unfold = - if eq_constr ty mkProp then + if eq_constr (fst evars) (EConstr.of_constr ty) EConstr.mkProp then (app_poly_sort prop env evars coq_all [| dom; lam |]), TypeGlobal.unfold_all else let forall = if prop then PropGlobal.coq_forall else TypeGlobal.coq_forall in @@ -1409,7 +1410,7 @@ module Strategies = let sigma = Sigma.Unsafe.of_evar_map (goalevars evars) in let Sigma (t', sigma, _) = rfn.Reductionops.e_redfun env sigma (EConstr.of_constr t) in let evars' = Sigma.to_evar_map sigma in - if eq_constr t' t then + if Termops.eq_constr evars' (EConstr.of_constr t') (EConstr.of_constr t) then state, Identity else state, Success { rew_car = ty; rew_from = t; rew_to = t'; @@ -1553,14 +1554,15 @@ let assert_replacing id newt tac = in let env' = Environ.reset_with_named_context (val_of_named_context nc) env in Refine.refine ~unsafe:false { run = begin fun sigma -> + let open EConstr in let Sigma (ev, sigma, p) = Evarutil.new_evar env' sigma (EConstr.of_constr concl) in let Sigma (ev', sigma, q) = Evarutil.new_evar env sigma (EConstr.of_constr newt) in let map d = let n = NamedDecl.get_id d in - if Id.equal n id then ev' else mkVar n + if Id.equal n id then ev' else EConstr.mkVar n in - let (e, _) = destEvar ev in - Sigma (EConstr.of_constr (mkEvar (e, Array.map_of_list map nc)), sigma, p +> q) + let (e, _) = EConstr.destEvar (Sigma.to_evar_map sigma) ev in + Sigma (mkEvar (e, Array.map_of_list map nc), sigma, p +> q) end } end } in Proofview.tclTHEN prf (Proofview.tclFOCUS 2 2 tac) @@ -1596,16 +1598,18 @@ let cl_rewrite_clause_newtac ?abs ?origsigma ~progress strat clause = convert_hyp_no_check (LocalAssum (id, newt)) <*> beta_hyp id | None, Some p -> + let p = EConstr.of_constr p in Proofview.Unsafe.tclEVARS undef <*> Proofview.Goal.enter { enter = begin fun gl -> let env = Proofview.Goal.env gl in let make = { run = begin fun sigma -> let Sigma (ev, sigma, q) = Evarutil.new_evar env sigma (EConstr.of_constr newt) in - Sigma (EConstr.of_constr (mkApp (p, [| ev |])), sigma, q) + Sigma (EConstr.mkApp (p, [| ev |]), sigma, q) end } in Refine.refine ~unsafe:false make <*> Proofview.Unsafe.tclNEWGOALS gls end } | None, None -> + let newt = EConstr.of_constr newt in Proofview.Unsafe.tclEVARS undef <*> convert_concl_no_check newt DEFAULTcast in @@ -2168,7 +2172,7 @@ let setoid_reflexivity = tac_open (poly_proof PropGlobal.get_reflexive_proof TypeGlobal.get_reflexive_proof env evm car rel) - (fun c -> tclCOMPLETE (apply c))) + (fun c -> tclCOMPLETE (apply (EConstr.of_constr c)))) (reflexivity_red true) let setoid_symmetry = @@ -2177,7 +2181,7 @@ let setoid_symmetry = tac_open (poly_proof PropGlobal.get_symmetric_proof TypeGlobal.get_symmetric_proof env evm car rel) - (fun c -> apply c)) + (fun c -> apply (EConstr.of_constr c))) (symmetry_red true) let setoid_transitivity c = @@ -2186,8 +2190,8 @@ let setoid_transitivity c = tac_open (poly_proof PropGlobal.get_transitive_proof TypeGlobal.get_transitive_proof env evm car rel) (fun proof -> match c with - | None -> eapply proof - | Some c -> apply_with_bindings (proof,ImplicitBindings [ c ]))) + | None -> eapply (EConstr.of_constr proof) + | Some c -> apply_with_bindings (EConstr.of_constr proof,ImplicitBindings [ c ]))) (transitivity_red true c) let setoid_symmetry_in id = @@ -2204,10 +2208,11 @@ let setoid_symmetry_in id = let he,c1,c2 = mkApp (equiv, Array.of_list others),c1,c2 in let new_hyp' = mkApp (he, [| c2 ; c1 |]) in let new_hyp = it_mkProd_or_LetIn new_hyp' binders in + let new_hyp = EConstr.of_constr new_hyp in Proofview.V82.of_tactic (tclTHENLAST (Tactics.assert_after_replacing id new_hyp) - (tclTHENLIST [ intros; setoid_symmetry; apply (mkVar id); Tactics.assumption ])) + (tclTHENLIST [ intros; setoid_symmetry; apply (EConstr.mkVar id); Tactics.assumption ])) gl) let _ = Hook.set Tactics.setoid_reflexivity setoid_reflexivity diff --git a/ltac/rewrite.mli b/ltac/rewrite.mli index 35c448351..bf56eec10 100644 --- a/ltac/rewrite.mli +++ b/ltac/rewrite.mli @@ -105,7 +105,7 @@ val setoid_symmetry_in : Id.t -> unit Proofview.tactic val setoid_reflexivity : unit Proofview.tactic -val setoid_transitivity : constr option -> unit Proofview.tactic +val setoid_transitivity : EConstr.constr option -> unit Proofview.tactic val apply_strategy : diff --git a/ltac/tacexpr.mli b/ltac/tacexpr.mli index 9c25a1645..b8d2d42b7 100644 --- a/ltac/tacexpr.mli +++ b/ltac/tacexpr.mli @@ -120,9 +120,9 @@ type glob_constr_pattern_and_expr = binding_bound_vars * glob_constr_and_expr * type 'a delayed_open = 'a Tactypes.delayed_open = { delayed : 'r. Environ.env -> 'r Sigma.t -> ('a, 'r) Sigma.sigma } -type delayed_open_constr_with_bindings = Term.constr with_bindings delayed_open +type delayed_open_constr_with_bindings = EConstr.constr with_bindings delayed_open -type delayed_open_constr = Term.constr delayed_open +type delayed_open_constr = EConstr.constr delayed_open type intro_pattern = delayed_open_constr intro_pattern_expr located type intro_patterns = delayed_open_constr intro_pattern_expr located list @@ -354,7 +354,7 @@ type raw_tactic_arg = (** Interpreted tactics *) -type t_trm = Term.constr +type t_trm = EConstr.constr type t_pat = constr_pattern type t_cst = evaluable_global_reference type t_ref = ltac_constant located diff --git a/ltac/tacinterp.ml b/ltac/tacinterp.ml index 142f061b5..553565639 100644 --- a/ltac/tacinterp.ml +++ b/ltac/tacinterp.ml @@ -746,11 +746,12 @@ let interp_closed_typed_pattern_with_occurrences ist env sigma (occs, a) = let interp_constr_with_occurrences_and_name_as_list = interp_constr_in_compound_list - (fun c -> ((AllOccurrences,c),Anonymous)) + (fun c -> ((AllOccurrences,EConstr.of_constr c),Anonymous)) (function ((occs,c),Anonymous) when occs == AllOccurrences -> c | _ -> raise Not_found) (fun ist env sigma (occ_c,na) -> let (sigma,c_interp) = interp_constr_with_occurrences ist env sigma occ_c in + let c_interp = (fst c_interp, EConstr.of_constr (snd c_interp)) in sigma, (c_interp, interp_name ist env sigma na)) @@ -853,7 +854,7 @@ let rec message_of_value v = Ftactic.return (int (out_gen (topwit wit_int) v)) else if has_type v (topwit wit_intro_pattern) then let p = out_gen (topwit wit_intro_pattern) v in - let print env sigma c = pr_constr_env env sigma (fst (Tactics.run_delayed env Evd.empty c)) in + let print env sigma c = pr_constr_env env sigma (EConstr.Unsafe.to_constr (fst (Tactics.run_delayed env Evd.empty c))) in Ftactic.nf_enter { enter = begin fun gl -> Ftactic.return (Miscprint.pr_intro_pattern (fun c -> print (pf_env gl) (project gl) c) p) end } @@ -917,6 +918,7 @@ and interp_intro_pattern_action ist env sigma = function let c = { delayed = fun env sigma -> let sigma = Sigma.to_evar_map sigma in let (sigma, c) = interp_open_constr ist env sigma c in + let c = EConstr.of_constr c in Sigma.Unsafe.of_pair (c, sigma) } in let sigma,ipat = interp_intro_pattern ist env sigma ipat in @@ -1002,6 +1004,8 @@ let interp_bindings ist env sigma = function let interp_constr_with_bindings ist env sigma (c,bl) = let sigma, bl = interp_bindings ist env sigma bl in let sigma, c = interp_open_constr ist env sigma c in + let c = EConstr.of_constr c in + let bl = Miscops.map_bindings EConstr.of_constr bl in sigma, (c,bl) let interp_open_constr_with_bindings ist env sigma (c,bl) = @@ -1021,6 +1025,7 @@ let interp_open_constr_with_bindings_loc ist ((c,_),bl as cb) = let f = { delayed = fun env sigma -> let sigma = Sigma.to_evar_map sigma in let (sigma, c) = interp_open_constr_with_bindings ist env sigma cb in + let c = Miscops.map_with_bindings EConstr.of_constr c in Sigma.Unsafe.of_pair (c, sigma) } in (loc,f) @@ -1044,7 +1049,7 @@ let interp_destruction_arg ist gl arg = then keep,ElimOnIdent (loc,id') else (keep, ElimOnConstr { delayed = begin fun env sigma -> - try Sigma.here (constr_of_id env id', NoBindings) sigma + try Sigma.here (EConstr.of_constr (constr_of_id env id'), NoBindings) sigma with Not_found -> user_err ~loc ~hdr:"interp_destruction_arg" ( pr_id id ++ strbrk " binds to " ++ pr_id id' ++ strbrk " which is neither a declared nor a quantified hypothesis.") @@ -1066,7 +1071,7 @@ let interp_destruction_arg ist gl arg = keep,ElimOnAnonHyp (out_gen (topwit wit_int) v) else match Value.to_constr v with | None -> error () - | Some c -> keep,ElimOnConstr { delayed = fun env sigma -> Sigma ((c,NoBindings), sigma, Sigma.refl) } + | Some c -> keep,ElimOnConstr { delayed = fun env sigma -> Sigma ((EConstr.of_constr c,NoBindings), sigma, Sigma.refl) } with Not_found -> (* We were in non strict (interactive) mode *) if Tactics.is_quantified_hypothesis id gl then @@ -1076,6 +1081,7 @@ let interp_destruction_arg ist gl arg = let f = { delayed = fun env sigma -> let sigma = Sigma.to_evar_map sigma in let (sigma,c) = interp_open_constr ist env sigma c in + let c = EConstr.of_constr c in Sigma.Unsafe.of_pair ((c,NoBindings), sigma) } in keep,ElimOnConstr f @@ -1701,7 +1707,7 @@ and interp_atomic ist tac : unit Proofview.tactic = let env = pf_env gl in let f sigma (id,n,c) = let (sigma,c_interp) = interp_type ist env sigma c in - sigma , (interp_ident ist env sigma id,n,c_interp) in + sigma , (interp_ident ist env sigma id,n,EConstr.of_constr c_interp) in let (sigma,l_interp) = Evd.MonadR.List.map_right (fun c sigma -> f sigma c) l (project gl) in @@ -1716,7 +1722,7 @@ and interp_atomic ist tac : unit Proofview.tactic = let env = pf_env gl in let f sigma (id,c) = let (sigma,c_interp) = interp_type ist env sigma c in - sigma , (interp_ident ist env sigma id,c_interp) in + sigma , (interp_ident ist env sigma id,EConstr.of_constr c_interp) in let (sigma,l_interp) = Evd.MonadR.List.map_right (fun c sigma -> f sigma c) l (project gl) in @@ -1731,6 +1737,7 @@ and interp_atomic ist tac : unit Proofview.tactic = let (sigma,c) = (if Option.is_empty t then interp_constr else interp_type) ist env sigma c in + let c = EConstr.of_constr c in let sigma, ipat' = interp_intro_pattern_option ist env sigma ipat in let tac = Option.map (Option.map (interp_tactic ist)) t in Tacticals.New.tclWITHHOLES false @@ -1758,6 +1765,7 @@ and interp_atomic ist tac : unit Proofview.tactic = if Locusops.is_nowhere clp then (* We try to fully-typecheck the term *) let (sigma,c_interp) = interp_constr ist env sigma c in + let c_interp = EConstr.of_constr c_interp in let let_tac b na c cl eqpat = let id = Option.default (Loc.ghost,IntroAnonymous) eqpat in let with_eq = if b then None else Some (true,id) in @@ -1776,11 +1784,12 @@ and interp_atomic ist tac : unit Proofview.tactic = Tactics.letin_pat_tac with_eq na c cl in let (sigma',c) = interp_pure_open_constr ist env sigma c in + let c = EConstr.of_constr c in name_atomic ~env (TacLetTac(na,c,clp,b,eqpat)) (Tacticals.New.tclWITHHOLES false (*in hope of a future "eset/epose"*) (let_pat_tac b (interp_name ist env sigma na) - ((sigma,sigma'),EConstr.of_constr c) clp eqpat) sigma') + ((sigma,sigma'),c) clp eqpat) sigma') end } (* Derived basic tactics *) @@ -1845,6 +1854,7 @@ and interp_atomic ist tac : unit Proofview.tactic = then interp_type ist (pf_env gl) sigma c else interp_constr ist (pf_env gl) sigma c in + let c = EConstr.of_constr c in Sigma.Unsafe.of_pair (c, sigma) end } in Tactics.change None c_interp (interp_clause ist (pf_env gl) (project gl) cl) @@ -1868,6 +1878,7 @@ and interp_atomic ist tac : unit Proofview.tactic = try let sigma = Sigma.to_evar_map sigma in let (sigma, c) = interp_constr ist env sigma c in + let c = EConstr.of_constr c in Sigma.Unsafe.of_pair (c, sigma) with e when to_catch e (* Hack *) -> user_err (strbrk "Failed to get enough information from the left-hand side to type the right-hand side.") @@ -1884,6 +1895,7 @@ and interp_atomic ist tac : unit Proofview.tactic = let f = { delayed = fun env sigma -> let sigma = Sigma.to_evar_map sigma in let (sigma, c) = interp_open_constr_with_bindings ist env sigma c in + let c = Miscops.map_with_bindings EConstr.of_constr c in Sigma.Unsafe.of_pair (c, sigma) } in (b,m,keep,f)) l in @@ -1906,6 +1918,7 @@ and interp_atomic ist tac : unit Proofview.tactic = | None -> sigma , None | Some c -> let (sigma,c_interp) = interp_constr ist env sigma c in + let c_interp = EConstr.of_constr c_interp in sigma , Some c_interp in let dqhyps = interp_declared_or_quantified_hypothesis ist env sigma hyp in @@ -1932,6 +1945,7 @@ and interp_atomic ist tac : unit Proofview.tactic = let env = Proofview.Goal.env gl in let sigma = project gl in let (sigma,c_interp) = interp_constr ist env sigma c in + let c_interp = EConstr.of_constr c_interp in let dqhyps = interp_declared_or_quantified_hypothesis ist env sigma hyp in let hyps = interp_hyp_list ist env sigma idl in let tac = name_atomic ~env @@ -2041,6 +2055,7 @@ end } let interp_bindings' ist bl = Ftactic.return { delayed = fun env sigma -> let (sigma, bl) = interp_bindings ist env (Sigma.to_evar_map sigma) bl in + let bl = Miscops.map_bindings EConstr.of_constr bl in Sigma.Unsafe.of_pair (bl, sigma) } diff --git a/ltac/tauto.ml b/ltac/tauto.ml index 11996af73..e3f5342de 100644 --- a/ltac/tauto.ml +++ b/ltac/tauto.ml @@ -161,8 +161,9 @@ let flatten_contravariant_conj _ ist = | Some (_,args) -> let args = List.map EConstr.Unsafe.to_constr args in let newtyp = List.fold_right mkArrow args c in + let newtyp = EConstr.of_constr newtyp in let intros = tclMAP (fun _ -> intro) args in - let by = tclTHENLIST [intros; apply hyp; split; assumption] in + let by = tclTHENLIST [intros; apply (EConstr.of_constr hyp); split; assumption] in tclTHENLIST [assert_ ~by newtyp; clear (destVar hyp)] | _ -> fail @@ -186,17 +187,17 @@ let flatten_contravariant_disj _ ist = let typ = assoc_var "X1" ist in let typ = EConstr.of_constr typ in let c = assoc_var "X2" ist in + let c = EConstr.of_constr c in let hyp = assoc_var "id" ist in match match_with_disjunction sigma ~strict:flags.strict_in_contravariant_hyp ~onlybinary:flags.binary_mode typ with | Some (_,args) -> - let args = List.map EConstr.Unsafe.to_constr args in let map i arg = - let typ = mkArrow arg c in + let typ = EConstr.mkArrow arg c in let ci = Tactics.constructor_tac false None (succ i) Misctypes.NoBindings in - let by = tclTHENLIST [intro; apply hyp; ci; assumption] in + let by = tclTHENLIST [intro; apply (EConstr.of_constr hyp); ci; assumption] in assert_ ~by typ in let tacs = List.mapi map args in @@ -231,6 +232,7 @@ let apply_nnpp _ ist = (Proofview.tclUNIT ()) begin fun () -> try let nnpp = Universes.constr_of_global (Nametab.global_of_path coq_nnpp_path) in + let nnpp = EConstr.of_constr nnpp in apply nnpp with Not_found -> tclFAIL 0 (Pp.mt ()) end diff --git a/plugins/btauto/refl_btauto.ml b/plugins/btauto/refl_btauto.ml index 1e49d8cad..27398cf65 100644 --- a/plugins/btauto/refl_btauto.ml +++ b/plugins/btauto/refl_btauto.ml @@ -249,9 +249,10 @@ module Btauto = struct let fl = reify env fl in let fr = reify env fr in let changed_gl = Term.mkApp (c, [|typ; fl; fr|]) in + let changed_gl = EConstr.of_constr changed_gl in Tacticals.New.tclTHENLIST [ Tactics.change_concl changed_gl; - Tactics.apply (Lazy.force soundness); + Tactics.apply (EConstr.of_constr (Lazy.force soundness)); Tactics.normalise_vm_in_concl; try_unification env ] diff --git a/plugins/cc/cctac.ml b/plugins/cc/cctac.ml index 7c78f3a17..7b023413d 100644 --- a/plugins/cc/cctac.ml +++ b/plugins/cc/cctac.ml @@ -238,17 +238,17 @@ let build_projection intype (cstr:pconstructor) special default gls= let _M =mkMeta let app_global f args k = - Tacticals.pf_constr_of_global (Lazy.force f) (fun fc -> k (mkApp (fc, args))) + Tacticals.pf_constr_of_global (Lazy.force f) (fun fc -> k (EConstr.of_constr (mkApp (fc, args)))) let new_app_global f args k = - Tacticals.New.pf_constr_of_global (Lazy.force f) (fun fc -> k (mkApp (fc, args))) + Tacticals.New.pf_constr_of_global (Lazy.force f) (fun fc -> k (EConstr.of_constr (mkApp (fc, args)))) -let new_refine c = Proofview.V82.tactic (refine (EConstr.of_constr c)) -let refine c = refine (EConstr.of_constr c) +let new_refine c = Proofview.V82.tactic (refine c) +let refine c = refine c let assert_before n c = Proofview.Goal.enter { enter = begin fun gl -> - let evm, _ = Tacmach.New.pf_apply type_of gl (EConstr.of_constr c) in + let evm, _ = Tacmach.New.pf_apply type_of gl c in Tacticals.New.tclTHEN (Proofview.V82.tactic (Refiner.tclEVARS evm)) (assert_before n c) end } @@ -269,7 +269,7 @@ let rec proof_tac p : unit Proofview.tactic = let type_of t = Tacmach.New.pf_unsafe_type_of gl (EConstr.of_constr t) in try (* type_of can raise exceptions *) match p.p_rule with - Ax c -> exact_check c + Ax c -> exact_check (EConstr.of_constr c) | SymAx c -> let l=constr_of_term p.p_lhs and r=constr_of_term p.p_rhs in @@ -333,6 +333,7 @@ let refute_tac c t1 t2 p = let tt1=constr_of_term t1 and tt2=constr_of_term t2 in let hid = Tacmach.New.of_old (pf_get_new_id (Id.of_string "Heq")) gl in let false_t=mkApp (c,[|mkVar hid|]) in + let false_t = EConstr.of_constr false_t in let k intype = let neweq= new_app_global _eq [|intype;tt1;tt2|] in Tacticals.New.tclTHENS (neweq (assert_before (Name hid))) @@ -341,7 +342,7 @@ let refute_tac c t1 t2 p = end } let refine_exact_check c gl = - let evm, _ = pf_apply type_of gl (EConstr.of_constr c) in + let evm, _ = pf_apply type_of gl c in Tacticals.tclTHEN (Refiner.tclEVARS evm) (Proofview.V82.of_tactic (exact_check c)) gl let convert_to_goal_tac c t1 t2 p = @@ -363,6 +364,8 @@ let convert_to_hyp_tac c1 t1 c2 t2 p = let tt2=constr_of_term t2 in let h = Tacmach.New.of_old (pf_get_new_id (Id.of_string "H")) gl in let false_t=mkApp (c2,[|mkVar h|]) in + let false_t = EConstr.of_constr false_t in + let tt2 = EConstr.of_constr tt2 in Tacticals.New.tclTHENS (assert_before (Name h) tt2) [convert_to_goal_tac c1 t1 t2 p; simplest_elim false_t] @@ -387,6 +390,7 @@ let discriminate_tac (cstr,u as cstru) p = [|intype;outtype;proj;t1;t2;mkVar hid|] in let endt k = injt (fun injt -> + let injt = EConstr.Unsafe.to_constr injt in app_global _eq_rect [|outtype;trivial;pred;identity;concl;injt|] k) in let neweq=new_app_global _eq [|intype;t1;t2|] in @@ -486,7 +490,7 @@ let mk_eq f c1 c2 k = let term = mkApp (fc, [| ty; c1; c2 |]) in let evm, _ = type_of (pf_env gl) evm (EConstr.of_constr term) in Tacticals.New.tclTHEN (Proofview.V82.tactic (Refiner.tclEVARS evm)) - (k term) + (k (EConstr.of_constr term)) end }) let f_equal = diff --git a/plugins/decl_mode/decl_proof_instr.ml b/plugins/decl_mode/decl_proof_instr.ml index 031a6253a..54206aa95 100644 --- a/plugins/decl_mode/decl_proof_instr.ml +++ b/plugins/decl_mode/decl_proof_instr.ml @@ -130,7 +130,7 @@ let clean_tmp gls = clean_all (tmp_ids gls) gls let assert_postpone id t = - assert_before (Name id) t + assert_before (Name id) (EConstr.of_constr t) (* start a proof *) @@ -268,6 +268,7 @@ let add_justification_hyps keep items gls = | _ -> let id=pf_get_new_id local_hyp_prefix gls in keep:=Id.Set.add id !keep; + let c = EConstr.of_constr c in tclTHEN (Proofview.V82.of_tactic (letin_tac None (Names.Name id) c None Locusops.nowhere)) (Proofview.V82.of_tactic (clear_body [id])) gls in tclMAP add_aux items gls @@ -488,6 +489,7 @@ let thus_tac c ctyp submetas gls = with Not_found -> error "I could not relate this statement to the thesis." in if List.is_empty list then + let proof = EConstr.of_constr proof in Proofview.V82.of_tactic (exact_check proof) gls else let refiner = concl_refiner list proof gls in @@ -546,7 +548,7 @@ let decompose_eq id gls = let whd = (special_whd gls typ) in match kind_of_term whd with App (f,args)-> - if eq_constr f (Lazy.force _eq) && Int.equal (Array.length args) 3 + if Term.eq_constr f (Lazy.force _eq) && Int.equal (Array.length args) 3 then (args.(0), args.(1), args.(2)) @@ -584,15 +586,15 @@ let instr_rew _thus rew_side cut gls0 = let new_eq = mkApp(Lazy.force _eq,[|typ;cut.cut_stat.st_it;rhs|]) in tclTHENS (Proofview.V82.of_tactic (assert_postpone c_id new_eq)) [tclTHEN tcl_erase_info - (tclTHENS (Proofview.V82.of_tactic (transitivity lhs)) - [just_tac;Proofview.V82.of_tactic (exact_check (mkVar last_id))]); + (tclTHENS (Proofview.V82.of_tactic (transitivity (EConstr.of_constr lhs))) + [just_tac;Proofview.V82.of_tactic (exact_check (EConstr.mkVar last_id))]); thus_tac new_eq] gls0 | Rhs -> let new_eq = mkApp(Lazy.force _eq,[|typ;lhs;cut.cut_stat.st_it|]) in tclTHENS (Proofview.V82.of_tactic (assert_postpone c_id new_eq)) [tclTHEN tcl_erase_info - (tclTHENS (Proofview.V82.of_tactic (transitivity rhs)) - [Proofview.V82.of_tactic (exact_check (mkVar last_id));just_tac]); + (tclTHENS (Proofview.V82.of_tactic (transitivity (EConstr.of_constr rhs))) + [Proofview.V82.of_tactic (exact_check (EConstr.mkVar last_id));just_tac]); thus_tac new_eq] gls0 @@ -772,7 +774,7 @@ let rec consider_match may_intro introduced available expected gls = try conjunction_arity id gls with Not_found -> error "Matching hypothesis not found." in tclTHENLIST - [Proofview.V82.of_tactic (simplest_case (mkVar id)); + [Proofview.V82.of_tactic (simplest_case (EConstr.mkVar id)); intron_then nhyps [] (fun l -> consider_match may_intro introduced (List.rev_append l rest_ids) expected)] gls) @@ -780,7 +782,8 @@ let rec consider_match may_intro introduced available expected gls = gls let consider_tac c hyps gls = - match kind_of_term (strip_outer_cast (project gls) (EConstr.of_constr c)) with + let c = EConstr.of_constr c in + match kind_of_term (strip_outer_cast (project gls) c) with Var id -> consider_match false [] [id] hyps gls | _ -> @@ -817,6 +820,7 @@ let rec build_function sigma args body = let define_tac id args body gls = let t = build_function (project gls) args body in + let t = EConstr.of_constr t in Proofview.V82.of_tactic (letin_tac None (Name id) t None Locusops.nowhere) gls (* tactics for reconsider *) @@ -828,6 +832,7 @@ let cast_tac id_or_thesis typ gls = | Thesis (For _ ) -> error "\"thesis for ...\" is not applicable here." | Thesis Plain -> + let typ = EConstr.of_constr typ in Proofview.V82.of_tactic (convert_concl typ DEFAULTcast) gls (* per cases *) @@ -1087,7 +1092,7 @@ let thesis_for obj typ per_info env= ((Printer.pr_constr_env env Evd.empty obj) ++ spc () ++ str"cannot give an induction hypothesis (wrong inductive type).") in let params,args = List.chop per_info.per_nparams all_args in - let _ = if not (List.for_all2 eq_constr params per_info.per_params) then + let _ = if not (List.for_all2 Term.eq_constr params per_info.per_params) then user_err ~hdr:"thesis_for" ((Printer.pr_constr_env env Evd.empty obj) ++ spc () ++ str "cannot give an induction hypothesis (wrong parameters).") in @@ -1219,10 +1224,10 @@ let hrec_for fix_id per_info gls obj_id = let ind,u = destInd cind in assert (eq_ind ind per_info.per_ind); let params,args= List.chop per_info.per_nparams all_args in assert begin - try List.for_all2 eq_constr params per_info.per_params with + try List.for_all2 Term.eq_constr params per_info.per_params with Invalid_argument _ -> false end; let hd2 = applist (mkVar fix_id,args@[obj]) in - compose_lam rc (Reductionops.whd_beta gls.sigma (EConstr.of_constr hd2)) + EConstr.of_constr (compose_lam rc (Reductionops.whd_beta gls.sigma (EConstr.of_constr hd2))) let warn_missing_case = CWarnings.create ~name:"declmode-missing-case" ~category:"declmode" @@ -1336,7 +1341,7 @@ let my_refine c gls = let oc = { run = begin fun sigma -> let sigma = Sigma.to_evar_map sigma in let (sigma, c) = understand_my_constr (pf_env gls) sigma c (pf_concl gls) in - Sigma.Unsafe.of_pair (c, sigma) + Sigma.Unsafe.of_pair (EConstr.of_constr c, sigma) end } in Proofview.V82.of_tactic (Tactics.New.refine oc) gls @@ -1366,14 +1371,14 @@ let end_tac et2 gls = begin match et,ek with _,EK_unknown -> - tclSOLVE [Proofview.V82.of_tactic (simplest_elim pi.per_casee)] + tclSOLVE [Proofview.V82.of_tactic (simplest_elim (EConstr.of_constr pi.per_casee))] | ET_Case_analysis,EK_nodep -> tclTHEN - (Proofview.V82.of_tactic (simplest_case pi.per_casee)) + (Proofview.V82.of_tactic (simplest_case (EConstr.of_constr pi.per_casee))) (default_justification (List.map mkVar clauses)) | ET_Induction,EK_nodep -> tclTHENLIST - [Proofview.V82.of_tactic (generalize (pi.per_args@[pi.per_casee])); + [Proofview.V82.of_tactic (generalize (List.map EConstr.of_constr (pi.per_args@[pi.per_casee]))); Proofview.V82.of_tactic (simple_induct (AnonHyp (succ (List.length pi.per_args)))); default_justification (List.map mkVar clauses)] | ET_Case_analysis,EK_dep tree -> @@ -1385,7 +1390,7 @@ let end_tac et2 gls = (initial_instance_stack clauses) [pi.per_casee] 0 tree | ET_Induction,EK_dep tree -> let nargs = (List.length pi.per_args) in - tclTHEN (Proofview.V82.of_tactic (generalize (pi.per_args@[pi.per_casee]))) + tclTHEN (Proofview.V82.of_tactic (generalize (List.map EConstr.of_constr (pi.per_args@[pi.per_casee])))) begin fun gls0 -> let fix_id = diff --git a/plugins/decl_mode/decl_proof_instr.mli b/plugins/decl_mode/decl_proof_instr.mli index 325969dad..ba196ff01 100644 --- a/plugins/decl_mode/decl_proof_instr.mli +++ b/plugins/decl_mode/decl_proof_instr.mli @@ -89,7 +89,7 @@ val push_arg : Term.constr -> val hrec_for: Id.t -> Decl_mode.per_info -> Proof_type.goal Tacmach.sigma -> - Id.t -> Term.constr + Id.t -> EConstr.constr val consider_match : bool -> diff --git a/plugins/firstorder/instances.ml b/plugins/firstorder/instances.ml index 6c245063c..a320b47aa 100644 --- a/plugins/firstorder/instances.ml +++ b/plugins/firstorder/instances.ml @@ -117,6 +117,7 @@ let mk_open_instance id idc gl m t= let nid=(fresh_id avoid var_id gl) in let evmap = Sigma.Unsafe.of_evar_map evmap in let Sigma ((c, _), evmap, _) = Evarutil.new_type_evar env evmap Evd.univ_flexible in + let c = EConstr.Unsafe.to_constr c in let evmap = Sigma.to_evar_map evmap in let decl = LocalAssum (Name nid, c) in aux (n-1) (nid::avoid) (Environ.push_rel decl env) evmap (decl::decls) in @@ -131,13 +132,13 @@ let left_instance_tac (inst,id) continue seq= if lookup (id,None) seq then tclFAIL 0 (Pp.str "already done") else - tclTHENS (Proofview.V82.of_tactic (cut dom)) + tclTHENS (Proofview.V82.of_tactic (cut (EConstr.of_constr dom))) [tclTHENLIST [Proofview.V82.of_tactic introf; pf_constr_of_global id (fun idc -> (fun gls-> Proofview.V82.of_tactic (generalize - [mkApp(idc, - [|mkVar (Tacmach.pf_nth_hyp_id gls 1)|])]) gls)); + [EConstr.of_constr (mkApp(idc, + [|mkVar (Tacmach.pf_nth_hyp_id gls 1)|]))]) gls)); Proofview.V82.of_tactic introf; tclSOLVE [wrap 1 false continue (deepen (record (id,None) seq))]]; @@ -154,14 +155,15 @@ let left_instance_tac (inst,id) continue seq= let gt= it_mkLambda_or_LetIn (mkApp(idc,[|ot|])) rc in + let gt = EConstr.of_constr gt in let evmap, _ = - try Typing.type_of (pf_env gl) evmap (EConstr.of_constr gt) + try Typing.type_of (pf_env gl) evmap gt with e when CErrors.noncritical e -> error "Untypable instance, maybe higher-order non-prenex quantification" in tclTHEN (Refiner.tclEVARS evmap) (Proofview.V82.of_tactic (generalize [gt])) gl) else pf_constr_of_global id (fun idc -> - Proofview.V82.of_tactic (generalize [mkApp(idc,[|t|])])) + Proofview.V82.of_tactic (generalize [EConstr.of_constr (mkApp(idc,[|t|]))])) in tclTHENLIST [special_generalize; @@ -172,16 +174,16 @@ let left_instance_tac (inst,id) continue seq= let right_instance_tac inst continue seq= match inst with Phantom dom -> - tclTHENS (Proofview.V82.of_tactic (cut dom)) + tclTHENS (Proofview.V82.of_tactic (cut (EConstr.of_constr dom))) [tclTHENLIST [Proofview.V82.of_tactic introf; (fun gls-> Proofview.V82.of_tactic (split (ImplicitBindings - [mkVar (Tacmach.pf_nth_hyp_id gls 1)])) gls); + [EConstr.mkVar (Tacmach.pf_nth_hyp_id gls 1)])) gls); tclSOLVE [wrap 0 true continue (deepen seq)]]; tclTRY (Proofview.V82.of_tactic assumption)] | Real ((0,t),_) -> - (tclTHEN (Proofview.V82.of_tactic (split (ImplicitBindings [t]))) + (tclTHEN (Proofview.V82.of_tactic (split (ImplicitBindings [EConstr.of_constr t]))) (tclSOLVE [wrap 0 true continue (deepen seq)])) | Real ((m,t),_) -> tclFAIL 0 (Pp.str "not implemented ... yet") diff --git a/plugins/firstorder/rules.ml b/plugins/firstorder/rules.ml index 1d107e9af..bed7a727f 100644 --- a/plugins/firstorder/rules.ml +++ b/plugins/firstorder/rules.ml @@ -59,7 +59,7 @@ let clear_global=function (* connection rules *) let axiom_tac t seq= - try pf_constr_of_global (find_left t seq) (fun c -> Proofview.V82.of_tactic (exact_no_check c)) + try pf_constr_of_global (find_left t seq) (fun c -> Proofview.V82.of_tactic (exact_no_check (EConstr.of_constr c))) with Not_found->tclFAIL 0 (Pp.str "No axiom link") let ll_atom_tac a backtrack id continue seq= @@ -68,7 +68,7 @@ let ll_atom_tac a backtrack id continue seq= tclTHENLIST [pf_constr_of_global (find_left a seq) (fun left -> pf_constr_of_global id (fun id -> - Proofview.V82.of_tactic (generalize [mkApp(id, [|left|])]))); + Proofview.V82.of_tactic (generalize [EConstr.of_constr (mkApp(id, [|left|]))]))); clear_global id; Proofview.V82.of_tactic intro] with Not_found->tclFAIL 0 (Pp.str "No link")) @@ -95,7 +95,7 @@ let left_and_tac ind backtrack id continue seq gls= let n=(construct_nhyps ind gls).(0) in tclIFTHENELSE (tclTHENLIST - [Proofview.V82.of_tactic (Tacticals.New.pf_constr_of_global id simplest_elim); + [Proofview.V82.of_tactic (Tacticals.New.pf_constr_of_global id (EConstr.of_constr %> simplest_elim)); clear_global id; tclDO n (Proofview.V82.of_tactic intro)]) (wrap n false continue seq) @@ -109,12 +109,12 @@ let left_or_tac ind backtrack id continue seq gls= tclDO n (Proofview.V82.of_tactic intro); wrap n false continue seq] in tclIFTHENSVELSE - (Proofview.V82.of_tactic (Tacticals.New.pf_constr_of_global id simplest_elim)) + (Proofview.V82.of_tactic (Tacticals.New.pf_constr_of_global id (EConstr.of_constr %> simplest_elim))) (Array.map f v) backtrack gls let left_false_tac id= - Proofview.V82.of_tactic (Tacticals.New.pf_constr_of_global id simplest_elim) + Proofview.V82.of_tactic (Tacticals.New.pf_constr_of_global id (EConstr.of_constr %> simplest_elim)) (* left arrow connective rules *) @@ -131,7 +131,7 @@ let ll_ind_tac (ind,u as indu) largs backtrack id continue seq gl= let vars=Array.init p (fun j->mkRel (p-j)) in let capply=mkApp ((lift p cstr),vars) in let head=mkApp ((lift p idc),[|capply|]) in - it_mkLambda_or_LetIn head rc in + EConstr.of_constr (it_mkLambda_or_LetIn head rc) in let lp=Array.length rcs in let newhyps idc =List.init lp (myterm idc) in tclIFTHENELSE @@ -143,16 +143,16 @@ let ll_ind_tac (ind,u as indu) largs backtrack id continue seq gl= let ll_arrow_tac a b c backtrack id continue seq= let cc=mkProd(Anonymous,a,(lift 1 b)) in - let d idc =mkLambda (Anonymous,b, - mkApp (idc, [|mkLambda (Anonymous,(lift 1 a),(mkRel 2))|])) in + let d idc =EConstr.of_constr (mkLambda (Anonymous,b, + mkApp (idc, [|mkLambda (Anonymous,(lift 1 a),(mkRel 2))|]))) in tclORELSE - (tclTHENS (Proofview.V82.of_tactic (cut c)) + (tclTHENS (Proofview.V82.of_tactic (cut (EConstr.of_constr c))) [tclTHENLIST [Proofview.V82.of_tactic introf; clear_global id; wrap 1 false continue seq]; - tclTHENS (Proofview.V82.of_tactic (cut cc)) - [pf_constr_of_global id (fun c -> Proofview.V82.of_tactic (exact_no_check c)); + tclTHENS (Proofview.V82.of_tactic (cut (EConstr.of_constr cc))) + [pf_constr_of_global id (fun c -> Proofview.V82.of_tactic (exact_no_check (EConstr.of_constr c))); tclTHENLIST [pf_constr_of_global id (fun idc -> Proofview.V82.of_tactic (generalize [d idc])); clear_global id; @@ -177,7 +177,7 @@ let forall_tac backtrack continue seq= let left_exists_tac ind backtrack id continue seq gls= let n=(construct_nhyps ind gls).(0) in tclIFTHENELSE - (Proofview.V82.of_tactic (Tacticals.New.pf_constr_of_global id simplest_elim)) + (Proofview.V82.of_tactic (Tacticals.New.pf_constr_of_global id (EConstr.of_constr %> simplest_elim))) (tclTHENLIST [clear_global id; tclDO n (Proofview.V82.of_tactic intro); (wrap (n-1) false continue seq)]) @@ -186,13 +186,14 @@ let left_exists_tac ind backtrack id continue seq gls= let ll_forall_tac prod backtrack id continue seq= tclORELSE - (tclTHENS (Proofview.V82.of_tactic (cut prod)) + (tclTHENS (Proofview.V82.of_tactic (cut (EConstr.of_constr prod))) [tclTHENLIST [Proofview.V82.of_tactic intro; pf_constr_of_global id (fun idc -> (fun gls-> let id0=pf_nth_hyp_id gls 1 in let term=mkApp(idc,[|mkVar(id0)|]) in + let term = EConstr.of_constr term in tclTHEN (Proofview.V82.of_tactic (generalize [term])) (Proofview.V82.of_tactic (clear [id0])) gls)); clear_global id; Proofview.V82.of_tactic intro; diff --git a/plugins/fourier/fourierR.ml b/plugins/fourier/fourierR.ml index a14ec8a2c..fa64b276c 100644 --- a/plugins/fourier/fourierR.ml +++ b/plugins/fourier/fourierR.ml @@ -281,6 +281,8 @@ let fourier_lineq lineq1 = (* Defined constants *) let get = Lazy.force +let cget = get +let eget c = EConstr.of_constr (Lazy.force c) let constant = Coqlib.gen_constant "Fourier" (* Standard library *) @@ -373,6 +375,7 @@ let rational_to_real x = (* preuve que 0 False *) let tac_zero_inf_false gl (n,d) = - if n=0 then (apply (get coq_Rnot_lt0)) + let get = eget in +if n=0 then (apply (get coq_Rnot_lt0)) else (Tacticals.New.tclTHEN (apply (get coq_Rle_not_lt)) (tac_zero_infeq_pos gl (-n,d))) @@ -408,6 +413,7 @@ let tac_zero_inf_false gl (n,d) = (* preuve que 0<=(-n)*(1/d) => False *) let tac_zero_infeq_false gl (n,d) = + let get = eget in (Tacticals.New.tclTHEN (apply (get coq_Rlt_not_le_frac_opp)) (tac_zero_inf_pos gl (-n,d))) ;; @@ -415,7 +421,8 @@ let tac_zero_infeq_false gl (n,d) = let exact = exact_check;; let tac_use h = - let tac = exact h.hname in + let get = eget in + let tac = exact (EConstr.of_constr h.hname) in match h.htype with "Rlt" -> tac |"Rle" -> tac @@ -470,6 +477,7 @@ let rec fourier () = try match (kind_of_term goal) with App (f,args) -> + let get = eget in (match (string_of_R_constr f) with "Rlt" -> (Tacticals.New.tclTHEN @@ -548,6 +556,7 @@ let rec fourier () = !t2 |] in let tc=rational_to_real cres in (* puis sa preuve *) + let get = eget in let tac1=ref (if h1.hstrict then (Tacticals.New.tclTHENS (apply (get coq_Rfourier_lt)) [tac_use h1; @@ -584,29 +593,29 @@ let rec fourier () = then tac_zero_inf_false gl (rational_to_fraction cres) else tac_zero_infeq_false gl (rational_to_fraction cres) in - tac:=(Tacticals.New.tclTHENS (cut ineq) + tac:=(Tacticals.New.tclTHENS (cut (EConstr.of_constr ineq)) [Tacticals.New.tclTHEN (change_concl - (mkAppL [| get coq_not; ineq|] - )) + (EConstr.of_constr (mkAppL [| cget coq_not; ineq|] + ))) (Tacticals.New.tclTHEN (apply (if sres then get coq_Rnot_lt_lt else get coq_Rnot_le_le)) (Tacticals.New.tclTHENS (Equality.replace - (mkAppL [|get coq_Rminus;!t2;!t1|] + (mkAppL [|cget coq_Rminus;!t2;!t1|] ) tc) [tac2; (Tacticals.New.tclTHENS (Equality.replace - (mkApp (get coq_Rinv, - [|get coq_R1|])) - (get coq_R1)) + (mkApp (cget coq_Rinv, + [|cget coq_R1|])) + (cget coq_R1)) (* en attendant Field, ça peut aider Ring de remplacer 1/1 par 1 ... *) [Tacticals.New.tclORELSE (* TODO : Ring.polynom []*) (Proofview.tclUNIT ()) (Proofview.tclUNIT ()); - Tacticals.New.pf_constr_of_global (get coq_sym_eqT) (fun symeq -> - (Tacticals.New.tclTHEN (apply symeq) + Tacticals.New.pf_constr_of_global (cget coq_sym_eqT) (fun symeq -> + (Tacticals.New.tclTHEN (apply (EConstr.of_constr symeq)) (apply (get coq_Rinv_1))))] ) diff --git a/plugins/funind/functional_principles_proofs.ml b/plugins/funind/functional_principles_proofs.ml index b674f40e9..503cafdd5 100644 --- a/plugins/funind/functional_principles_proofs.ml +++ b/plugins/funind/functional_principles_proofs.ml @@ -175,6 +175,7 @@ let is_incompatible_eq t = res let change_hyp_with_using msg hyp_id t tac : tactic = + let t = EConstr.of_constr t in fun g -> let prov_id = pf_get_new_id hyp_id g in tclTHENS @@ -451,6 +452,7 @@ let clean_hyp_with_heq ptes_infos eq_hyps hyp_id env sigma = ) in let to_refine = EConstr.of_constr to_refine in + let t_x = EConstr.of_constr t_x in (* observe_tac "rec hyp " *) (tclTHENS (Proofview.V82.of_tactic (assert_before (Name rec_pte_id) t_x)) @@ -644,7 +646,8 @@ let instanciate_hyps_with_args (do_prove:Id.t list -> tactic) hyps args_id = fun g -> let prov_hid = pf_get_new_id hid g in let c = mkApp(mkVar hid,args) in - let evm, _ = pf_apply Typing.type_of g (EConstr.of_constr c) in + let c = EConstr.of_constr c in + let evm, _ = pf_apply Typing.type_of g c in tclTHENLIST[ Refiner.tclEVARS evm; Proofview.V82.of_tactic (pose_proof (Name prov_hid) c); @@ -709,13 +712,14 @@ let build_proof let term_eq = make_refl_eq (Lazy.force refl_equal) type_of_term t in + let term_eq = EConstr.of_constr term_eq in tclTHENSEQ [ - Proofview.V82.of_tactic (generalize (term_eq::(List.map mkVar dyn_infos.rec_hyps))); + Proofview.V82.of_tactic (generalize (term_eq::(List.map EConstr.mkVar dyn_infos.rec_hyps))); thin dyn_infos.rec_hyps; - Proofview.V82.of_tactic (pattern_option [Locus.AllOccurrencesBut [1],t] None); + Proofview.V82.of_tactic (pattern_option [Locus.AllOccurrencesBut [1],EConstr.of_constr t] None); (fun g -> observe_tac "toto" ( - tclTHENSEQ [Proofview.V82.of_tactic (Simple.case t); + tclTHENSEQ [Proofview.V82.of_tactic (Simple.case (EConstr.of_constr t)); (fun g' -> let g'_nb_prod = nb_prod (project g') (EConstr.of_constr (pf_concl g')) in let nb_instanciate_partial = g'_nb_prod - g_nb_prod in @@ -942,7 +946,7 @@ let generalize_non_dep hyp g = in (* observe (str "to_revert := " ++ prlist_with_sep spc Ppconstr.pr_id to_revert); *) tclTHEN - ((* observe_tac "h_generalize" *) (Proofview.V82.of_tactic (generalize (List.map mkVar to_revert) ))) + ((* observe_tac "h_generalize" *) (Proofview.V82.of_tactic (generalize (List.map EConstr.mkVar to_revert) ))) ((* observe_tac "thin" *) (thin to_revert)) g @@ -950,7 +954,7 @@ let id_of_decl = RelDecl.get_name %> Nameops.out_name let var_of_decl = id_of_decl %> mkVar let revert idl = tclTHEN - (Proofview.V82.of_tactic (generalize (List.map mkVar idl))) + (Proofview.V82.of_tactic (generalize (List.map EConstr.mkVar idl))) (thin idl) let generate_equation_lemma evd fnames f fun_num nb_params nb_args rec_args_num = @@ -991,7 +995,7 @@ let generate_equation_lemma evd fnames f fun_num nb_params nb_args rec_args_num let rec_id = pf_nth_hyp_id g 1 in tclTHENSEQ [observe_tac "generalize_non_dep in generate_equation_lemma" (generalize_non_dep rec_id); - observe_tac "h_case" (Proofview.V82.of_tactic (simplest_case (mkVar rec_id))); + observe_tac "h_case" (Proofview.V82.of_tactic (simplest_case (EConstr.mkVar rec_id))); (Proofview.V82.of_tactic intros_reflexivity)] g ) ] @@ -1064,10 +1068,11 @@ let do_replace (evd:Evd.evar_map ref) params rec_arg_num rev_args_id f fun_num a let prove_princ_for_struct (evd:Evd.evar_map ref) interactive_proof fun_num fnames all_funs _nparams : tactic = fun g -> let princ_type = pf_concl g in + let princ_type = EConstr.of_constr princ_type in (* Pp.msgnl (str "princ_type " ++ Printer.pr_lconstr princ_type); *) (* Pp.msgnl (str "all_funs "); *) (* Array.iter (fun c -> Pp.msgnl (Printer.pr_lconstr c)) all_funs; *) - let princ_info = compute_elim_sig princ_type in + let princ_info = compute_elim_sig (project g) princ_type in let fresh_id = let avoid = ref (pf_ids_of_hyps g) in (fun na -> @@ -1227,7 +1232,7 @@ let prove_princ_for_struct (evd:Evd.evar_map ref) interactive_proof fun_num fnam | _, this_fix_info::others_infos -> let other_fix_infos = List.map - (fun fi -> fi.name,fi.idx + 1 ,fi.types) + (fun fi -> fi.name,fi.idx + 1 ,EConstr.of_constr fi.types) (pre_info@others_infos) in if List.is_empty other_fix_infos @@ -1462,11 +1467,11 @@ let new_prove_with_tcc is_mes acc_inv hrec tcc_hyps eqs : tactic = backtrack_eqs_until_hrec hrec eqs; (* observe_tac ("new_prove_with_tcc ( applying "^(Id.to_string hrec)^" )" ) *) (tclTHENS (* We must have exactly ONE subgoal !*) - (Proofview.V82.of_tactic (apply (mkVar hrec))) + (Proofview.V82.of_tactic (apply (EConstr.mkVar hrec))) [ tclTHENSEQ [ (Proofview.V82.of_tactic (keep (tcc_hyps@eqs))); - (Proofview.V82.of_tactic (apply (Lazy.force acc_inv))); + (Proofview.V82.of_tactic (apply (EConstr.of_constr (Lazy.force acc_inv)))); (fun g -> if is_mes then @@ -1482,7 +1487,7 @@ let new_prove_with_tcc is_mes acc_inv hrec tcc_hyps eqs : tactic = tclCOMPLETE( Eauto.eauto_with_bases (true,5) - [{ Tacexpr.delayed = fun _ sigma -> Sigma.here (Lazy.force refl_equal) sigma}] + [{ Tacexpr.delayed = fun _ sigma -> Sigma.here (EConstr.of_constr (Lazy.force refl_equal)) sigma}] [Hints.Hint_db.empty empty_transparent_state false] ) ) @@ -1518,7 +1523,8 @@ let prove_principle_for_gen (f_ref,functional_ref,eq_ref) tcc_lemma_ref is_mes rec_arg_num rec_arg_type relation gl = let princ_type = pf_concl gl in - let princ_info = compute_elim_sig princ_type in + let princ_type = EConstr.of_constr princ_type in + let princ_info = compute_elim_sig (project gl) princ_type in let fresh_id = let avoid = ref (pf_ids_of_hyps gl) in fun na -> @@ -1572,7 +1578,7 @@ let prove_principle_for_gen Nameops.out_name (fresh_id (Name (Id.of_string ("Acc_"^(Id.to_string rec_arg_id))))) in let revert l = - tclTHEN (Proofview.V82.of_tactic (Tactics.generalize (List.map mkVar l))) (Proofview.V82.of_tactic (clear l)) + tclTHEN (Proofview.V82.of_tactic (Tactics.generalize (List.map EConstr.mkVar l))) (Proofview.V82.of_tactic (clear l)) in let fix_id = Nameops.out_name (fresh_id (Name hrec_id)) in let prove_rec_arg_acc g = @@ -1580,12 +1586,12 @@ let prove_principle_for_gen (tclCOMPLETE (tclTHEN (Proofview.V82.of_tactic (assert_by (Name wf_thm_id) - (mkApp (delayed_force well_founded,[|input_type;relation|])) + (EConstr.of_constr (mkApp (delayed_force well_founded,[|input_type;relation|]))) (Proofview.V82.tactic (fun g -> (* observe_tac "prove wf" *) (tclCOMPLETE (wf_tac is_mes)) g)))) ( (* observe_tac *) (* "apply wf_thm" *) - Proofview.V82.of_tactic (Tactics.Simple.apply (mkApp(mkVar wf_thm_id,[|mkVar rec_arg_id|]))) + Proofview.V82.of_tactic (Tactics.Simple.apply (EConstr.of_constr (mkApp(mkVar wf_thm_id,[|mkVar rec_arg_id|])))) ) ) ) @@ -1596,7 +1602,7 @@ let prove_principle_for_gen let lemma = match !tcc_lemma_ref with | None -> error "No tcc proof !!" - | Some lemma -> lemma + | Some lemma -> EConstr.of_constr lemma in (* let rec list_diff del_list check_list = *) (* match del_list with *) @@ -1644,7 +1650,7 @@ let prove_principle_for_gen ); (* observe_tac "" *) Proofview.V82.of_tactic (assert_by (Name acc_rec_arg_id) - (mkApp (delayed_force acc_rel,[|input_type;relation;mkVar rec_arg_id|])) + (EConstr.of_constr (mkApp (delayed_force acc_rel,[|input_type;relation;mkVar rec_arg_id|]))) (Proofview.V82.tactic prove_rec_arg_acc) ); (* observe_tac "reverting" *) (revert (List.rev (acc_rec_arg_id::args_ids))); diff --git a/plugins/funind/functional_principles_types.ml b/plugins/funind/functional_principles_types.ml index 4b47b83af..4d88f9f91 100644 --- a/plugins/funind/functional_principles_types.ml +++ b/plugins/funind/functional_principles_types.ml @@ -28,7 +28,8 @@ let observe s = a functional one *) let compute_new_princ_type_from_rel rel_to_fun sorts princ_type = - let princ_type_info = compute_elim_sig princ_type in + let princ_type = EConstr.of_constr princ_type in + let princ_type_info = compute_elim_sig Evd.empty princ_type (** FIXME *) in let env = Global.env () in let env_with_params = Environ.push_rel_context princ_type_info.params env in let tbl = Hashtbl.create 792 in @@ -89,7 +90,7 @@ let compute_new_princ_type_from_rel rel_to_fun sorts princ_type = (Option.fold_right mkProd_or_LetIn princ_type_info.indarg - princ_type_info.concl + (EConstr.Unsafe.to_constr princ_type_info.concl) ) princ_type_info.args ) @@ -243,7 +244,8 @@ let compute_new_princ_type_from_rel rel_to_fun sorts princ_type = let change_property_sort evd toSort princ princName = let open Context.Rel.Declaration in - let princ_info = compute_elim_sig princ in + let princ = EConstr.of_constr princ in + let princ_info = compute_elim_sig evd princ in let change_sort_in_predicate decl = LocalAssum (get_name decl, @@ -270,7 +272,7 @@ let change_property_sort evd toSort princ princName = let build_functional_principle (evd:Evd.evar_map ref) interactive_proof old_princ_type sorts funs i proof_tac hook = (* First we get the type of the old graph principle *) - let mutr_nparams = (compute_elim_sig old_princ_type).nparams in + let mutr_nparams = (compute_elim_sig !evd (EConstr.of_constr old_princ_type)).nparams in (* let time1 = System.get_time () in *) let new_principle_type = compute_new_princ_type_from_rel diff --git a/plugins/funind/g_indfun.ml4 b/plugins/funind/g_indfun.ml4 index 6603a95a8..a6f971703 100644 --- a/plugins/funind/g_indfun.ml4 +++ b/plugins/funind/g_indfun.ml4 @@ -58,7 +58,7 @@ let pr_fun_ind_using_typed prc prlc _ opt_c = | None -> mt () | Some b -> let (b, _) = Tactics.run_delayed (Global.env ()) Evd.empty b in - spc () ++ hov 2 (str "using" ++ spc () ++ pr_with_bindings_typed prc prlc b) + spc () ++ hov 2 (str "using" ++ spc () ++ pr_with_bindings_typed (EConstr.Unsafe.to_constr %> prc) (EConstr.Unsafe.to_constr %> prlc) b) ARGUMENT EXTEND fun_ind_using @@ -97,7 +97,9 @@ ARGUMENT EXTEND with_names TYPED AS intropattern_opt PRINTED BY pr_intro_as_pat | [] ->[ None ] END - +let functional_induction b c x pat = + let x = Option.map (Miscops.map_with_bindings EConstr.Unsafe.to_constr) x in + Proofview.V82.tactic (functional_induction true c x (Option.map out_disjunctive pat)) TACTIC EXTEND newfunind @@ -108,7 +110,7 @@ TACTIC EXTEND newfunind | [c] -> c | c::cl -> applist(c,cl) in - Extratactics.onSomeWithHoles (fun x -> Proofview.V82.tactic (functional_induction true c x (Option.map out_disjunctive pat))) princl ] + Extratactics.onSomeWithHoles (fun x -> functional_induction true c x pat) princl ] END (***** debug only ***) TACTIC EXTEND snewfunind @@ -119,7 +121,7 @@ TACTIC EXTEND snewfunind | [c] -> c | c::cl -> applist(c,cl) in - Extratactics.onSomeWithHoles (fun x -> Proofview.V82.tactic (functional_induction false c x (Option.map out_disjunctive pat))) princl ] + Extratactics.onSomeWithHoles (fun x -> functional_induction false c x pat) princl ] END diff --git a/plugins/funind/indfun.ml b/plugins/funind/indfun.ml index e3ba52246..37a76bec1 100644 --- a/plugins/funind/indfun.ml +++ b/plugins/funind/indfun.ml @@ -81,7 +81,8 @@ let functional_induction with_clean c princl pat = | Some ((princ,binding)) -> princ,binding,Tacmach.pf_unsafe_type_of g (EConstr.of_constr princ),g in - let princ_infos = Tactics.compute_elim_sig princ_type in + let princ_type = EConstr.of_constr princ_type in + let princ_infos = Tactics.compute_elim_sig (Tacmach.project g') princ_type in let args_as_induction_constr = let c_list = if princ_infos.Tactics.farg_in_concl @@ -89,9 +90,11 @@ let functional_induction with_clean c princl pat = in let encoded_pat_as_patlist = List.make (List.length args + List.length c_list - 1) None @ [pat] in - List.map2 (fun c pat -> ((None,Tacexpr.ElimOnConstr ({ Tacexpr.delayed = fun env sigma -> Sigma ((c,NoBindings), sigma, Sigma.refl) })),(None,pat),None)) + List.map2 (fun c pat -> ((None,Tacexpr.ElimOnConstr ({ Tacexpr.delayed = fun env sigma -> Sigma ((EConstr.of_constr c,NoBindings), sigma, Sigma.refl) })),(None,pat),None)) (args@c_list) encoded_pat_as_patlist in + let princ = EConstr.of_constr princ in + let bindings = Miscops.map_bindings EConstr.of_constr bindings in let princ' = Some (princ,bindings) in let princ_vars = List.fold_right diff --git a/plugins/funind/invfun.ml b/plugins/funind/invfun.ml index b2419b1a5..36fb6dad3 100644 --- a/plugins/funind/invfun.ml +++ b/plugins/funind/invfun.ml @@ -252,7 +252,8 @@ let prove_fun_correct evd functional_induction funs_constr graphs_constr schemes (* and the principle to use in this lemma in $\zeta$ normal form *) let f_principle,princ_type = schemes.(i) in let princ_type = nf_zeta (EConstr.of_constr princ_type) in - let princ_infos = Tactics.compute_elim_sig princ_type in + let princ_type = EConstr.of_constr princ_type in + let princ_infos = Tactics.compute_elim_sig evd princ_type in (* The number of args of the function is then easily computable *) let nb_fun_args = nb_prod (project g) (EConstr.of_constr (pf_concl g)) - 2 in let args_names = generate_fresh_id (Id.of_string "x") [] nb_fun_args in @@ -315,7 +316,7 @@ let prove_fun_correct evd functional_induction funs_constr graphs_constr schemes match kind_of_term t'',kind_of_term t''' with | App(eq,args), App(graph',_) when - (eq_constr eq eq_ind) && + (Term.eq_constr eq eq_ind) && Array.exists (Constr.eq_constr_nounivs graph') graphs_constr -> (args.(2)::(mkApp(mkVar hid,[|args.(2);(mkApp(eq_construct,[|args.(0);args.(2)|]))|])) ::acc) @@ -387,7 +388,7 @@ let prove_fun_correct evd functional_induction funs_constr graphs_constr schemes observe_tac "rewriting res value" (Proofview.V82.of_tactic (Equality.rewriteLR (mkVar hres))); (* Conclusion *) observe_tac "exact" (fun g -> - Proofview.V82.of_tactic (exact_check (app_constructor g)) g) + Proofview.V82.of_tactic (exact_check (EConstr.of_constr (app_constructor g))) g) ] ) g @@ -440,7 +441,7 @@ let prove_fun_correct evd functional_induction funs_constr graphs_constr schemes observe_tac "principle" (Proofview.V82.of_tactic (assert_by (Name principle_id) princ_type - (exact_check f_principle))); + (exact_check (EConstr.of_constr f_principle)))); observe_tac "intro args_names" (tclMAP (fun id -> Proofview.V82.of_tactic (Simple.intro id)) args_names); (* observe_tac "titi" (pose_proof (Name (Id.of_string "__")) (Reductionops.nf_beta Evd.empty ((mkApp (mkVar principle_id,Array.of_list bindings))))); *) observe_tac "idtac" tclIDTAC; @@ -450,7 +451,7 @@ let prove_fun_correct evd functional_induction funs_constr graphs_constr schemes (fun gl -> let term = mkApp (mkVar principle_id,Array.of_list bindings) in let gl', _ty = pf_eapply (Typing.type_of ~refresh:true) gl (EConstr.of_constr term) in - Proofview.V82.of_tactic (apply term) gl') + Proofview.V82.of_tactic (apply (EConstr.of_constr term)) gl') )) (fun i g -> observe_tac ("proving branche "^string_of_int i) (prove_branche i) g ) ] @@ -467,7 +468,7 @@ let generalize_dependent_of x hyp g = tclMAP (function | LocalAssum (id,t) when not (Id.equal id hyp) && - (Termops.occur_var (pf_env g) (project g) x (EConstr.of_constr t)) -> tclTHEN (Proofview.V82.of_tactic (Tactics.generalize [mkVar id])) (thin [id]) + (Termops.occur_var (pf_env g) (project g) x (EConstr.of_constr t)) -> tclTHEN (Proofview.V82.of_tactic (Tactics.generalize [EConstr.mkVar id])) (thin [id]) | _ -> tclIDTAC ) (pf_hyps g) @@ -495,7 +496,7 @@ and intros_with_rewrite_aux : tactic = | Prod(_,t,t') -> begin match kind_of_term t with - | App(eq,args) when (eq_constr eq eq_ind) -> + | App(eq,args) when (Term.eq_constr eq eq_ind) -> if Reductionops.is_conv (pf_env g) (project g) (EConstr.of_constr args.(1)) (EConstr.of_constr args.(2)) then let id = pf_get_new_id (Id.of_string "y") g in @@ -541,11 +542,11 @@ and intros_with_rewrite_aux : tactic = intros_with_rewrite ] g end - | Ind _ when eq_constr t (Coqlib.build_coq_False ()) -> + | Ind _ when Term.eq_constr t (Coqlib.build_coq_False ()) -> Proofview.V82.of_tactic tauto g | Case(_,_,v,_) -> tclTHENSEQ[ - Proofview.V82.of_tactic (simplest_case v); + Proofview.V82.of_tactic (simplest_case (EConstr.of_constr v)); intros_with_rewrite ] g | LetIn _ -> @@ -582,7 +583,7 @@ let rec reflexivity_with_destruct_cases g = match kind_of_term (snd (destApp (pf_concl g))).(2) with | Case(_,_,v,_) -> tclTHENSEQ[ - Proofview.V82.of_tactic (simplest_case v); + Proofview.V82.of_tactic (simplest_case (EConstr.of_constr v)); Proofview.V82.of_tactic intros; observe_tac "reflexivity_with_destruct_cases" reflexivity_with_destruct_cases ] @@ -597,7 +598,7 @@ let rec reflexivity_with_destruct_cases g = None -> tclIDTAC g | Some id -> match kind_of_term (pf_unsafe_type_of g (EConstr.mkVar id)) with - | App(eq,[|_;t1;t2|]) when eq_constr eq eq_ind -> + | App(eq,[|_;t1;t2|]) when Term.eq_constr eq eq_ind -> if Equality.discriminable (pf_env g) (project g) (EConstr.of_constr t1) (EConstr.of_constr t2) then Proofview.V82.of_tactic (Equality.discrHyp id) g else if Equality.injectable (pf_env g) (project g) (EConstr.of_constr t1) (EConstr.of_constr t2) @@ -662,7 +663,8 @@ let prove_fun_complete funcs graphs schemes lemmas_types_infos i : tactic = let f = funcs.(i) in let graph_principle = nf_zeta (EConstr.of_constr schemes.(i)) in let princ_type = pf_unsafe_type_of g (EConstr.of_constr graph_principle) in - let princ_infos = Tactics.compute_elim_sig princ_type in + let princ_type = EConstr.of_constr princ_type in + let princ_infos = Tactics.compute_elim_sig (project g) princ_type in (* Then we get the number of argument of the function and compute a fresh name for each of them *) @@ -717,7 +719,7 @@ let prove_fun_complete funcs graphs schemes lemmas_types_infos i : tactic = }) Locusops.onConcl) ; - Proofview.V82.of_tactic (generalize (List.map mkVar ids)); + Proofview.V82.of_tactic (generalize (List.map EConstr.mkVar ids)); thin ids ] else @@ -756,10 +758,10 @@ let prove_fun_complete funcs graphs schemes lemmas_types_infos i : tactic = tclTHENSEQ [ tclMAP (fun id -> Proofview.V82.of_tactic (Simple.intro id)) (args_names@[res;hres]); observe_tac "h_generalize" - (Proofview.V82.of_tactic (generalize [mkApp(applist(graph_principle,params),Array.map (fun c -> applist(c,params)) lemmas)])); + (Proofview.V82.of_tactic (generalize [EConstr.of_constr (mkApp(applist(graph_principle,params),Array.map (fun c -> applist(c,params)) lemmas))])); Proofview.V82.of_tactic (Simple.intro graph_principle_id); observe_tac "" (tclTHEN_i - (observe_tac "elim" (Proofview.V82.of_tactic (elim false None (mkVar hres,NoBindings) (Some (mkVar graph_principle_id,NoBindings))))) + (observe_tac "elim" (Proofview.V82.of_tactic (elim false None (EConstr.mkVar hres,NoBindings) (Some (EConstr.mkVar graph_principle_id,NoBindings))))) (fun i g -> observe_tac "prove_branche" (prove_branche i) g )) ] g @@ -939,7 +941,7 @@ let revert_graph kn post_tac hid g = let f_args,res = Array.chop (Array.length args - 1) args in tclTHENSEQ [ - Proofview.V82.of_tactic (generalize [applist(mkConst f_complete,(Array.to_list f_args)@[res.(0);mkVar hid])]); + Proofview.V82.of_tactic (generalize [EConstr.of_constr (applist(mkConst f_complete,(Array.to_list f_args)@[res.(0);mkVar hid]))]); thin [hid]; Proofview.V82.of_tactic (Simple.intro hid); post_tac hid @@ -972,18 +974,18 @@ let functional_inversion kn hid fconst f_correct : tactic = let old_ids = List.fold_right Id.Set.add (pf_ids_of_hyps g) Id.Set.empty in let type_of_h = pf_unsafe_type_of g (EConstr.mkVar hid) in match kind_of_term type_of_h with - | App(eq,args) when eq_constr eq (make_eq ()) -> + | App(eq,args) when Term.eq_constr eq (make_eq ()) -> let pre_tac,f_args,res = match kind_of_term args.(1),kind_of_term args.(2) with - | App(f,f_args),_ when eq_constr f fconst -> + | App(f,f_args),_ when Term.eq_constr f fconst -> ((fun hid -> Proofview.V82.of_tactic (intros_symmetry (Locusops.onHyp hid))),f_args,args.(2)) - |_,App(f,f_args) when eq_constr f fconst -> + |_,App(f,f_args) when Term.eq_constr f fconst -> ((fun hid -> tclIDTAC),f_args,args.(1)) | _ -> (fun hid -> tclFAIL 1 (mt ())),[||],args.(2) in tclTHENSEQ[ pre_tac hid; - Proofview.V82.of_tactic (generalize [applist(f_correct,(Array.to_list f_args)@[res;mkVar hid])]); + Proofview.V82.of_tactic (generalize [EConstr.of_constr (applist(f_correct,(Array.to_list f_args)@[res;mkVar hid]))]); thin [hid]; Proofview.V82.of_tactic (Simple.intro hid); Proofview.V82.of_tactic (Inv.inv FullInversion None (NamedHyp hid)); @@ -1024,7 +1026,7 @@ let invfun qhyp f g = (fun hid -> Proofview.V82.tactic begin fun g -> let hyp_typ = pf_unsafe_type_of g (EConstr.mkVar hid) in match kind_of_term hyp_typ with - | App(eq,args) when eq_constr eq (make_eq ()) -> + | App(eq,args) when Term.eq_constr eq (make_eq ()) -> begin let f1,_ = decompose_app args.(1) in try diff --git a/plugins/funind/merge.ml b/plugins/funind/merge.ml index 222c0c804..3688b8c15 100644 --- a/plugins/funind/merge.ml +++ b/plugins/funind/merge.ml @@ -32,7 +32,7 @@ module RelDecl = Context.Rel.Declaration (** {2 Useful operations on constr and glob_constr} *) -let rec popn i c = if i<=0 then c else pop (EConstr.of_constr (popn (i-1) c)) +let rec popn i c = if i<=0 then c else EConstr.of_constr (pop (popn (i-1) c)) (** Substitutions in constr *) let compare_constr_nosub t1 t2 = @@ -979,19 +979,20 @@ let funify_branches relinfo nfuns branch = let relprinctype_to_funprinctype relprinctype nfuns = - let relinfo = compute_elim_sig relprinctype in + let relprinctype = EConstr.of_constr relprinctype in + let relinfo = compute_elim_sig Evd.empty (** FIXME*) relprinctype in assert (not relinfo.farg_in_concl); assert (relinfo.indarg_in_concl); (* first remove indarg and indarg_in_concl *) let relinfo_noindarg = { relinfo with indarg_in_concl = false; indarg = None; - concl = remove_last_arg (pop (EConstr.of_constr relinfo.concl)); } in + concl = EConstr.of_constr (remove_last_arg (pop relinfo.concl)); } in (* the nfuns last induction arguments are functional ones: remove them *) let relinfo_argsok = { relinfo_noindarg with nargs = relinfo_noindarg.nargs - nfuns; (* args is in reverse order, so remove fst *) args = remove_n_fst_list nfuns relinfo_noindarg.args; - concl = popn nfuns relinfo_noindarg.concl + concl = popn nfuns relinfo_noindarg.concl; } in let new_branches = List.map (funify_branches relinfo_argsok nfuns) relinfo_argsok.branches in diff --git a/plugins/funind/recdef.ml b/plugins/funind/recdef.ml index b2c93a754..d5ee42af8 100644 --- a/plugins/funind/recdef.ml +++ b/plugins/funind/recdef.ml @@ -286,7 +286,7 @@ let tclUSER tac is_mes l g = let tclUSER_if_not_mes concl_tac is_mes names_to_suppress = if is_mes - then tclCOMPLETE (fun gl -> Proofview.V82.of_tactic (Simple.apply (delayed_force well_founded_ltof)) gl) + then tclCOMPLETE (fun gl -> Proofview.V82.of_tactic (Simple.apply (EConstr.of_constr (delayed_force well_founded_ltof))) gl) else (* tclTHEN (Simple.apply (delayed_force acc_intro_generator_function) ) *) (tclUSER concl_tac is_mes names_to_suppress) @@ -465,7 +465,7 @@ let rec travel_aux jinfo continuation_tac (expr_info:constr infos) = end | App _ -> let f,args = decompose_app expr_info.info in - if eq_constr f (expr_info.f_constr) + if Term.eq_constr f (expr_info.f_constr) then jinfo.app_reC (f,args) expr_info continuation_tac expr_info else begin @@ -517,21 +517,21 @@ let rec prove_lt hyple g = let h = List.find (fun id -> match decompose_app (pf_unsafe_type_of g (EConstr.mkVar id)) with - | _, t::_ -> eq_constr t varx + | _, t::_ -> Term.eq_constr t varx | _ -> false ) hyple in let y = List.hd (List.tl (snd (decompose_app (pf_unsafe_type_of g (EConstr.mkVar h))))) in observe_tclTHENLIST (str "prove_lt1")[ - Proofview.V82.of_tactic (apply (mkApp(le_lt_trans (),[|varx;y;varz;mkVar h|]))); + Proofview.V82.of_tactic (apply (EConstr.of_constr (mkApp(le_lt_trans (),[|varx;y;varz;mkVar h|])))); observe_tac (str "prove_lt") (prove_lt hyple) ] with Not_found -> ( ( observe_tclTHENLIST (str "prove_lt2")[ - Proofview.V82.of_tactic (apply (delayed_force lt_S_n)); + Proofview.V82.of_tactic (apply (EConstr.of_constr (delayed_force lt_S_n))); (observe_tac (str "assumption: " ++ Printer.pr_goal g) (Proofview.V82.of_tactic assumption)) ]) ) @@ -549,15 +549,15 @@ let rec destruct_bounds_aux infos (bound,hyple,rechyps) lbounds g = let ids = h'::ids in let def = next_ident_away_in_goal def_id ids in observe_tclTHENLIST (str "destruct_bounds_aux1")[ - Proofview.V82.of_tactic (split (ImplicitBindings [s_max])); + Proofview.V82.of_tactic (split (ImplicitBindings [EConstr.of_constr s_max])); Proofview.V82.of_tactic (intro_then (fun id -> Proofview.V82.tactic begin observe_tac (str "destruct_bounds_aux") - (tclTHENS (Proofview.V82.of_tactic (simplest_case (mkVar id))) + (tclTHENS (Proofview.V82.of_tactic (simplest_case (EConstr.mkVar id))) [ observe_tclTHENLIST (str "")[Proofview.V82.of_tactic (intro_using h_id); - Proofview.V82.of_tactic (simplest_elim(mkApp(delayed_force lt_n_O,[|s_max|]))); + Proofview.V82.of_tactic (simplest_elim(EConstr.of_constr (mkApp(delayed_force lt_n_O,[|s_max|])))); Proofview.V82.of_tactic default_full_auto]; observe_tclTHENLIST (str "destruct_bounds_aux2")[ observe_tac (str "clearing k ") (Proofview.V82.of_tactic (clear [id])); @@ -588,7 +588,7 @@ let rec destruct_bounds_aux infos (bound,hyple,rechyps) lbounds g = ] g | (_,v_bound)::l -> observe_tclTHENLIST (str "destruct_bounds_aux3")[ - Proofview.V82.of_tactic (simplest_elim (mkVar v_bound)); + Proofview.V82.of_tactic (simplest_elim (EConstr.mkVar v_bound)); Proofview.V82.of_tactic (clear [v_bound]); tclDO 2 (Proofview.V82.of_tactic intro); onNthHypId 1 @@ -597,7 +597,7 @@ let rec destruct_bounds_aux infos (bound,hyple,rechyps) lbounds g = (fun p -> observe_tclTHENLIST (str "destruct_bounds_aux4")[ Proofview.V82.of_tactic (simplest_elim - (mkApp(delayed_force max_constr, [| bound; mkVar p|]))); + (EConstr.of_constr (mkApp(delayed_force max_constr, [| bound; mkVar p|])))); tclDO 3 (Proofview.V82.of_tactic intro); onNLastHypsId 3 (fun lids -> match lids with @@ -622,7 +622,7 @@ let terminate_app f_and_args expr_info continuation_tac infos = observe_tclTHENLIST (str "terminate_app1")[ continuation_tac infos; observe_tac (str "first split") - (Proofview.V82.of_tactic (split (ImplicitBindings [infos.info]))); + (Proofview.V82.of_tactic (split (ImplicitBindings [EConstr.of_constr infos.info]))); observe_tac (str "destruct_bounds (1)") (destruct_bounds infos) ] else continuation_tac infos @@ -633,7 +633,7 @@ let terminate_others _ expr_info continuation_tac infos = observe_tclTHENLIST (str "terminate_others")[ continuation_tac infos; observe_tac (str "first split") - (Proofview.V82.of_tactic (split (ImplicitBindings [infos.info]))); + (Proofview.V82.of_tactic (split (ImplicitBindings [EConstr.of_constr infos.info]))); observe_tac (str "destruct_bounds") (destruct_bounds infos) ] else continuation_tac infos @@ -657,7 +657,7 @@ let terminate_letin (na,b,t,e) expr_info continuation_tac info = continuation_tac {info with info = new_e; forbidden_ids = new_forbidden} let pf_type c tac gl = - let evars, ty = Typing.type_of (pf_env gl) (project gl) (EConstr.of_constr c) in + let evars, ty = Typing.type_of (pf_env gl) (project gl) c in tclTHEN (Refiner.tclEVARS evars) (tac ty) gl let pf_typel l tac = @@ -687,16 +687,18 @@ let mkDestructEq : let type_of_expr = pf_unsafe_type_of g (EConstr.of_constr expr) in let new_hyps = mkApp(Lazy.force refl_equal, [|type_of_expr; expr|]):: to_revert_constr in + let new_hyps = List.map EConstr.of_constr new_hyps in pf_typel new_hyps (fun _ -> observe_tclTHENLIST (str "mkDestructEq") [Proofview.V82.of_tactic (generalize new_hyps); (fun g2 -> let changefun patvars = { run = fun sigma -> let redfun = pattern_occs [Locus.AllOccurrencesBut [1], EConstr.of_constr expr] in - redfun.Reductionops.e_redfun (pf_env g2) sigma (EConstr.of_constr (pf_concl g2)) + let Sigma (c, sigma, p) = redfun.Reductionops.e_redfun (pf_env g2) sigma (EConstr.of_constr (pf_concl g2)) in + Sigma (EConstr.of_constr c, sigma, p) } in Proofview.V82.of_tactic (change_in_concl None changefun) g2); - Proofview.V82.of_tactic (simplest_case expr)]), to_revert + Proofview.V82.of_tactic (simplest_case (EConstr.of_constr expr))]), to_revert let terminate_case next_step (ci,a,t,l) expr_info continuation_tac infos g = @@ -742,7 +744,7 @@ let terminate_app_rec (f,args) expr_info continuation_tac _ = then observe_tclTHENLIST (str "terminate_app_rec1")[ observe_tac (str "first split") - (Proofview.V82.of_tactic (split (ImplicitBindings [new_infos.info]))); + (Proofview.V82.of_tactic (split (ImplicitBindings [EConstr.of_constr new_infos.info]))); observe_tac (str "destruct_bounds (3)") (destruct_bounds new_infos) ] @@ -751,7 +753,7 @@ let terminate_app_rec (f,args) expr_info continuation_tac _ = ] with Not_found -> observe_tac (str "terminate_app_rec not found") (tclTHENS - (Proofview.V82.of_tactic (simplest_elim (mkApp(mkVar expr_info.ih,Array.of_list args)))) + (Proofview.V82.of_tactic (simplest_elim (EConstr.of_constr (mkApp(mkVar expr_info.ih,Array.of_list args))))) [ observe_tclTHENLIST (str "terminate_app_rec2")[ Proofview.V82.of_tactic (intro_using rec_res_id); @@ -772,7 +774,7 @@ let terminate_app_rec (f,args) expr_info continuation_tac _ = then observe_tclTHENLIST (str "terminate_app_rec4")[ observe_tac (str "first split") - (Proofview.V82.of_tactic (split (ImplicitBindings [new_infos.info]))); + (Proofview.V82.of_tactic (split (ImplicitBindings [EConstr.of_constr new_infos.info]))); observe_tac (str "destruct_bounds (2)") (destruct_bounds new_infos) ] @@ -785,7 +787,7 @@ let terminate_app_rec (f,args) expr_info continuation_tac _ = ]; observe_tac (str "proving decreasing") ( tclTHENS (* proof of args < formal args *) - (Proofview.V82.of_tactic (apply (Lazy.force expr_info.acc_inv))) + (Proofview.V82.of_tactic (apply (EConstr.of_constr (Lazy.force expr_info.acc_inv)))) [ observe_tac (str "assumption") (Proofview.V82.of_tactic assumption); observe_tclTHENLIST (str "terminate_app_rec5") @@ -833,7 +835,7 @@ let rec prove_le g = in tclFIRST[ Proofview.V82.of_tactic assumption; - Proofview.V82.of_tactic (apply (delayed_force le_n)); + Proofview.V82.of_tactic (apply (EConstr.of_constr (delayed_force le_n))); begin try let matching_fun = @@ -846,7 +848,7 @@ let rec prove_le g = List.hd (List.tl args) in observe_tclTHENLIST (str "prove_le")[ - Proofview.V82.of_tactic (apply(mkApp(le_trans (),[|x;y;z;mkVar h|]))); + Proofview.V82.of_tactic (apply(EConstr.of_constr (mkApp(le_trans (),[|x;y;z;mkVar h|])))); observe_tac (str "prove_le (rec)") (prove_le) ] with Not_found -> tclFAIL 0 (mt()) @@ -876,7 +878,7 @@ let rec make_rewrite_list expr_info max = function ) [make_rewrite_list expr_info max l; observe_tclTHENLIST (str "make_rewrite_list")[ (* x < S max proof *) - Proofview.V82.of_tactic (apply (delayed_force le_lt_n_Sm)); + Proofview.V82.of_tactic (apply (EConstr.of_constr (delayed_force le_lt_n_Sm))); observe_tac (str "prove_le(2)") prove_le ] ] ) @@ -916,7 +918,7 @@ let make_rewrite expr_info l hp max = ])) ; observe_tclTHENLIST (str "make_rewrite1")[ (* x < S (S max) proof *) - Proofview.V82.of_tactic (apply (delayed_force le_lt_SS)); + Proofview.V82.of_tactic (apply (EConstr.of_constr (delayed_force le_lt_SS))); observe_tac (str "prove_le (3)") prove_le ] ]) @@ -928,7 +930,7 @@ let rec compute_max rew_tac max l = | (_,p,_)::l -> observe_tclTHENLIST (str "compute_max")[ Proofview.V82.of_tactic (simplest_elim - (mkApp(delayed_force max_constr, [| max; mkVar p|]))); + (EConstr.of_constr (mkApp(delayed_force max_constr, [| max; mkVar p|])))); tclDO 3 (Proofview.V82.of_tactic intro); onNLastHypsId 3 (fun lids -> match lids with @@ -947,7 +949,7 @@ let rec destruct_hex expr_info acc l = end | (v,hex)::l -> observe_tclTHENLIST (str "destruct_hex")[ - Proofview.V82.of_tactic (simplest_case (mkVar hex)); + Proofview.V82.of_tactic (simplest_case (EConstr.mkVar hex)); Proofview.V82.of_tactic (clear [hex]); tclDO 2 (Proofview.V82.of_tactic intro); onNthHypId 1 (fun hp -> @@ -995,13 +997,13 @@ let equation_app_rec (f,args) expr_info continuation_tac info = if expr_info.is_final && expr_info.is_main_branch then observe_tclTHENLIST (str "equation_app_rec") - [ Proofview.V82.of_tactic (simplest_case (mkApp (expr_info.f_terminate,Array.of_list args))); + [ Proofview.V82.of_tactic (simplest_case (EConstr.of_constr (mkApp (expr_info.f_terminate,Array.of_list args)))); continuation_tac {expr_info with args_assoc = (args,delayed_force coq_O)::expr_info.args_assoc}; observe_tac (str "app_rec intros_values_eq") (intros_values_eq expr_info []) ] else observe_tclTHENLIST (str "equation_app_rec1")[ - Proofview.V82.of_tactic (simplest_case (mkApp (expr_info.f_terminate,Array.of_list args))); + Proofview.V82.of_tactic (simplest_case (EConstr.of_constr (mkApp (expr_info.f_terminate,Array.of_list args)))); observe_tac (str "app_rec not_found") (continuation_tac {expr_info with args_assoc = (args,delayed_force coq_O)::expr_info.args_assoc}) ] end @@ -1086,9 +1088,9 @@ let termination_proof_header is_mes input_type ids args_id relation (str "first assert") (Proofview.V82.of_tactic (assert_before (Name wf_rec_arg) - (mkApp (delayed_force acc_rel, + (EConstr.of_constr (mkApp (delayed_force acc_rel, [|input_type;relation;mkVar rec_arg_id|]) - ) + )) )) ) [ @@ -1098,7 +1100,7 @@ let termination_proof_header is_mes input_type ids args_id relation (str "second assert") (Proofview.V82.of_tactic (assert_before (Name wf_thm) - (mkApp (delayed_force well_founded,[|input_type;relation|])) + (EConstr.of_constr (mkApp (delayed_force well_founded,[|input_type;relation|]))) )) ) [ @@ -1107,7 +1109,7 @@ let termination_proof_header is_mes input_type ids args_id relation (* this gives the accessibility argument *) observe_tac (str "apply wf_thm") - (Proofview.V82.of_tactic (Simple.apply (mkApp(mkVar wf_thm,[|mkVar rec_arg_id|]))) + (Proofview.V82.of_tactic (Simple.apply (EConstr.of_constr (mkApp(mkVar wf_thm,[|mkVar rec_arg_id|])))) ) ] ; @@ -1116,7 +1118,7 @@ let termination_proof_header is_mes input_type ids args_id relation [observe_tac (str "generalize") (onNLastHypsId (nargs+1) (tclMAP (fun id -> - tclTHEN (Proofview.V82.of_tactic (Tactics.generalize [mkVar id])) (Proofview.V82.of_tactic (clear [id]))) + tclTHEN (Proofview.V82.of_tactic (Tactics.generalize [EConstr.mkVar id])) (Proofview.V82.of_tactic (clear [id]))) )) ; observe_tac (str "fix") (Proofview.V82.of_tactic (fix (Some hrec) (nargs+1))); @@ -1214,7 +1216,7 @@ let build_and_l l = | Prod(_,_,t') -> is_well_founded t' | App(_,_) -> let (f,_) = decompose_app t in - eq_constr f (well_founded ()) + Term.eq_constr f (well_founded ()) | _ -> false in @@ -1231,7 +1233,7 @@ let build_and_l l = let c,tac,nb = f pl in mk_and p1 c, tclTHENS - (Proofview.V82.of_tactic (apply (constr_of_global conj_constr))) + (Proofview.V82.of_tactic (apply (EConstr.of_constr (constr_of_global conj_constr)))) [tclIDTAC; tac ],nb+1 @@ -1297,6 +1299,7 @@ let open_new_goal build_proof sigma using_lemmas ref_ goal_name (gls_type,decomp in let lemma = mkConst (Names.Constant.make1 (Lib.make_kn na)) in ref_ := Some lemma ; + let lemma = EConstr.of_constr lemma in let lid = ref [] in let h_num = ref (-1) in let env = Global.env () in @@ -1323,7 +1326,7 @@ let open_new_goal build_proof sigma using_lemmas ref_ goal_name (gls_type,decomp ] gls) (fun g -> match kind_of_term (pf_concl g) with - | App(f,_) when eq_constr f (well_founded ()) -> + | App(f,_) when Term.eq_constr f (well_founded ()) -> Proofview.V82.of_tactic (Auto.h_auto None [] (Some [])) g | _ -> incr h_num; @@ -1332,11 +1335,11 @@ let open_new_goal build_proof sigma using_lemmas ref_ goal_name (gls_type,decomp tclCOMPLETE( tclFIRST[ tclTHEN - (Proofview.V82.of_tactic (eapply_with_bindings (mkVar (List.nth !lid !h_num), NoBindings))) + (Proofview.V82.of_tactic (eapply_with_bindings (EConstr.mkVar (List.nth !lid !h_num), NoBindings))) (Proofview.V82.of_tactic e_assumption); Eauto.eauto_with_bases (true,5) - [{ Tacexpr.delayed = fun _ sigma -> Sigma.here (Lazy.force refl_equal) sigma}] + [{ Tacexpr.delayed = fun _ sigma -> Sigma.here (EConstr.of_constr (Lazy.force refl_equal)) sigma}] [Hints.Hint_db.empty empty_transparent_state false] ] ) @@ -1366,7 +1369,7 @@ let open_new_goal build_proof sigma using_lemmas ref_ goal_name (gls_type,decomp (fun c -> Proofview.V82.of_tactic (Tacticals.New.tclTHENLIST [intros; - Simple.apply (fst (interp_constr (Global.env()) Evd.empty c)) (*FIXME*); + Simple.apply (EConstr.of_constr (fst (interp_constr (Global.env()) Evd.empty c))) (*FIXME*); Tacticals.New.tclCOMPLETE Auto.default_auto ]) ) @@ -1428,8 +1431,8 @@ let start_equation (f:global_reference) (term_f:global_reference) h_intros x; Proofview.V82.of_tactic (unfold_in_concl [(Locus.AllOccurrences, evaluable_of_global_reference f)]); observe_tac (str "simplest_case") - (Proofview.V82.of_tactic (simplest_case (mkApp (terminate_constr, - Array.of_list (List.map mkVar x))))); + (Proofview.V82.of_tactic (simplest_case (EConstr.of_constr (mkApp (terminate_constr, + Array.of_list (List.map mkVar x)))))); observe_tac (str "prove_eq") (cont_tactic x)]) g;; let (com_eqn : int -> Id.t -> diff --git a/plugins/micromega/coq_micromega.ml b/plugins/micromega/coq_micromega.ml index a943ef2b0..f96b189c5 100644 --- a/plugins/micromega/coq_micromega.ml +++ b/plugins/micromega/coq_micromega.ml @@ -1461,7 +1461,7 @@ let rec make_goal_of_formula dexpr form = xset (Term.mkNamedLetIn (Names.Id.of_string name) expr typ acc) l in - xset concl l + EConstr.of_constr (xset concl l) end (** * MODULE END: M @@ -2000,12 +2000,12 @@ let micromega_gen (Tacticals.New.tclTHEN tac_arith tac)) in Tacticals.New.tclTHENS - (Tactics.forward true (Some None) (ipat_of_name goal_name) arith_goal) + (Tactics.forward true (Some None) (ipat_of_name goal_name) (EConstr.of_constr arith_goal)) [ kill_arith; (Tacticals.New.tclTHENLIST - [(Tactics.generalize (List.map Term.mkVar ids)); - Tactics.exact_check (Term.applist (Term.mkVar goal_name, arith_args)) + [(Tactics.generalize (List.map EConstr.mkVar ids)); + Tactics.exact_check (EConstr.of_constr (Term.applist (Term.mkVar goal_name, arith_args))) ] ) ] with @@ -2114,12 +2114,12 @@ let micromega_genr prover tac = (Tacticals.New.tclTHEN tac_arith tac)) in Tacticals.New.tclTHENS - (Tactics.forward true (Some None) (ipat_of_name goal_name) arith_goal) + (Tactics.forward true (Some None) (ipat_of_name goal_name) (EConstr.of_constr arith_goal)) [ kill_arith; (Tacticals.New.tclTHENLIST - [(Tactics.generalize (List.map Term.mkVar ids)); - Tactics.exact_check (Term.applist (Term.mkVar goal_name, arith_args)) + [(Tactics.generalize (List.map EConstr.mkVar ids)); + Tactics.exact_check (EConstr.of_constr (Term.applist (Term.mkVar goal_name, arith_args))) ] ) ] diff --git a/plugins/nsatz/nsatz.ml b/plugins/nsatz/nsatz.ml index 36bce780b..cc0c4277e 100644 --- a/plugins/nsatz/nsatz.ml +++ b/plugins/nsatz/nsatz.ml @@ -625,6 +625,7 @@ let nsatz lpol = let return_term t = let a = mkApp(gen_constant "CC" ["Init";"Logic"] "refl_equal",[|tllp ();t|]) in + let a = EConstr.of_constr a in generalize [a] let nsatz_compute t = diff --git a/plugins/omega/coq_omega.ml b/plugins/omega/coq_omega.ml index b832250a5..35d763ccc 100644 --- a/plugins/omega/coq_omega.ml +++ b/plugins/omega/coq_omega.ml @@ -38,9 +38,9 @@ open OmegaSolver let elim_id id = Proofview.Goal.nf_enter { enter = begin fun gl -> - simplest_elim (Tacmach.New.pf_global id gl) + simplest_elim (EConstr.of_constr (Tacmach.New.pf_global id gl)) end } -let resolve_id id gl = Proofview.V82.of_tactic (apply (pf_global gl id)) gl +let resolve_id id gl = Proofview.V82.of_tactic (apply (EConstr.of_constr (pf_global gl id))) gl let timing timer_name f arg = f arg @@ -149,7 +149,7 @@ let mk_then = tclTHENLIST let exists_tac c = constructor_tac false (Some 1) 1 (ImplicitBindings [c]) -let generalize_tac t = generalize t +let generalize_tac t = generalize (List.map EConstr.of_constr t) let elim t = simplest_elim t let exact t = Tacmach.refine t let unfold s = Tactics.unfold_in_concl [Locus.AllOccurrences, Lazy.force s] @@ -373,7 +373,7 @@ let mk_minus t1 t2 = mkApp (Lazy.force coq_Zminus, [| t1;t2 |]) let mk_eq t1 t2 = mkApp (Universes.constr_of_global (build_coq_eq ()), [| Lazy.force coq_Z; t1; t2 |]) let mk_le t1 t2 = mkApp (Lazy.force coq_Zle, [| t1; t2 |]) -let mk_gt t1 t2 = mkApp (Lazy.force coq_Zgt, [| t1; t2 |]) +let mk_gt t1 t2 = EConstr.of_constr (mkApp (Lazy.force coq_Zgt, [| t1; t2 |])) let mk_inv t = mkApp (Lazy.force coq_Zopp, [| t |]) let mk_and t1 t2 = mkApp (build_coq_and (), [| t1; t2 |]) let mk_or t1 t2 = mkApp (build_coq_or (), [| t1; t2 |]) @@ -569,6 +569,7 @@ let abstract_path typ path t = let focused_simpl path gl = let newc = context (fun i t -> pf_nf gl (EConstr.of_constr t)) (List.rev path) (pf_concl gl) in + let newc = EConstr.of_constr newc in Proofview.V82.of_tactic (convert_concl_no_check newc DEFAULTcast) gl let focused_simpl path = focused_simpl path @@ -1116,7 +1117,7 @@ let replay_history tactic_normalisation = let state_eg = mk_eq eq1 rhs in let tac = scalar_norm_add [P_APP 3] e2.body in Tacticals.New.tclTHENS - (cut state_eg) + (cut (EConstr.of_constr state_eg)) [ Tacticals.New.tclTHENS (Tacticals.New.tclTHENLIST [ (intros_using [aux]); @@ -1185,7 +1186,7 @@ let replay_history tactic_normalisation = if e1.kind == DISE then let tac = scalar_norm [P_APP 3] e2.body in Tacticals.New.tclTHENS - (cut state_eq) + (cut (EConstr.of_constr state_eq)) [Tacticals.New.tclTHENLIST [ (intros_using [aux1]); (generalize_tac @@ -1197,7 +1198,7 @@ let replay_history tactic_normalisation = Tacticals.New.tclTHEN (Proofview.V82.tactic (mk_then tac)) reflexivity ] else let tac = scalar_norm [P_APP 3] e2.body in - Tacticals.New.tclTHENS (cut state_eq) + Tacticals.New.tclTHENS (cut (EConstr.of_constr state_eq)) [ Tacticals.New.tclTHENS (cut (mk_gt kk izero)) @@ -1227,7 +1228,7 @@ let replay_history tactic_normalisation = scalar_norm [P_APP 3] e1.body in Tacticals.New.tclTHENS - (cut (mk_eq eq1 (mk_inv eq2))) + (cut (EConstr.of_constr (mk_eq eq1 (mk_inv eq2)))) [Tacticals.New.tclTHENLIST [ (intros_using [aux]); (generalize_tac [mkApp (Lazy.force coq_OMEGA8, @@ -1260,7 +1261,7 @@ let replay_history tactic_normalisation = shuffle_mult_right p_initial orig.body m ({c= negone;v= v}::def.body) in Tacticals.New.tclTHENS - (cut theorem) + (cut (EConstr.of_constr theorem)) [Tacticals.New.tclTHENLIST [ (intros_using [aux]); (elim_id aux); @@ -1273,7 +1274,7 @@ let replay_history tactic_normalisation = (clear [aux]); (intros_using [id]); (loop l) ]; - Tacticals.New.tclTHEN (exists_tac eq1) reflexivity ] + Tacticals.New.tclTHEN (exists_tac (EConstr.of_constr eq1)) reflexivity ] | SPLIT_INEQ(e,(e1,act1),(e2,act2)) :: l -> let id1 = new_identifier () and id2 = new_identifier () in @@ -1283,7 +1284,7 @@ let replay_history tactic_normalisation = let tac2 = scalar_norm_add [P_APP 2;P_TYPE] e.body in let eq = val_of(decompile e) in Tacticals.New.tclTHENS - (simplest_elim (applist (Lazy.force coq_OMEGA19, [eq; mkVar id]))) + (simplest_elim (EConstr.of_constr (applist (Lazy.force coq_OMEGA19, [eq; mkVar id])))) [Tacticals.New.tclTHENLIST [ Proofview.V82.tactic (mk_then tac1); (intros_using [id1]); (loop act1) ]; Tacticals.New.tclTHENLIST [ Proofview.V82.tactic (mk_then tac2); (intros_using [id2]); (loop act2) ]] | SUM(e3,(k1,e1),(k2,e2)) :: l -> @@ -1433,7 +1434,7 @@ let coq_omega = let i = new_id () in tag_hypothesis id i; (Tacticals.New.tclTHENLIST [ - (simplest_elim (applist (Lazy.force coq_intro_Z, [t]))); + (simplest_elim (EConstr.of_constr (applist (Lazy.force coq_intro_Z, [t])))); (intros_using [v; id]); (elim_id id); (clear [id]); @@ -1444,7 +1445,7 @@ let coq_omega = constant = zero; id = i} :: sys else (Tacticals.New.tclTHENLIST [ - (simplest_elim (applist (Lazy.force coq_new_var, [t]))); + (simplest_elim (EConstr.of_constr (applist (Lazy.force coq_new_var, [t])))); (intros_using [v;th]); tac ]), sys) @@ -1494,7 +1495,7 @@ let nat_inject = let id = new_identifier () in Tacticals.New.tclTHENS (Tacticals.New.tclTHEN - (simplest_elim (applist (Lazy.force coq_le_gt_dec, [t2;t1]))) + (simplest_elim (EConstr.of_constr (applist (Lazy.force coq_le_gt_dec, [t2;t1])))) (intros_using [id])) [ Tacticals.New.tclTHENLIST [ @@ -1793,15 +1794,15 @@ let destructure_hyps = | Kapp(Nat,_) -> Tacticals.New.tclTHENLIST [ (simplest_elim - (mkApp - (Lazy.force coq_not_eq, [|t1;t2;mkVar i|]))); + (EConstr.of_constr (mkApp + (Lazy.force coq_not_eq, [|t1;t2;mkVar i|])))); (onClearedName i (fun _ -> loop lit)) ] | Kapp(Z,_) -> Tacticals.New.tclTHENLIST [ (simplest_elim - (mkApp - (Lazy.force coq_not_Zeq, [|t1;t2;mkVar i|]))); + (EConstr.of_constr (mkApp + (Lazy.force coq_not_Zeq, [|t1;t2;mkVar i|])))); (onClearedName i (fun _ -> loop lit)) ] | _ -> loop lit @@ -1851,7 +1852,7 @@ let destructure_goal = (Proofview.V82.tactic (Tacmach.refine (EConstr.of_constr (mkApp (Lazy.force coq_dec_not_not, [| t; dec; mkNewMeta () |]))))) intro - with Undecidable -> Tactics.elim_type (build_coq_False ()) + with Undecidable -> Tactics.elim_type (EConstr.of_constr (build_coq_False ())) in Tacticals.New.tclTHEN goal_tac destructure_hyps in diff --git a/plugins/quote/quote.ml b/plugins/quote/quote.ml index 7b6d502b5..2cc402e28 100644 --- a/plugins/quote/quote.ml +++ b/plugins/quote/quote.ml @@ -458,8 +458,8 @@ let quote f lid = | _ -> assert false in match ivs.variable_lhs with - | None -> Tactics.convert_concl (mkApp (f, [| p |])) DEFAULTcast - | Some _ -> Tactics.convert_concl (mkApp (f, [| vm; p |])) DEFAULTcast + | None -> Tactics.convert_concl (EConstr.of_constr (mkApp (f, [| p |]))) DEFAULTcast + | Some _ -> Tactics.convert_concl (EConstr.of_constr (mkApp (f, [| vm; p |]))) DEFAULTcast end } let gen_quote cont c f lid = diff --git a/plugins/romega/refl_omega.ml b/plugins/romega/refl_omega.ml index ba882e39a..ab5033601 100644 --- a/plugins/romega/refl_omega.ml +++ b/plugins/romega/refl_omega.ml @@ -1222,7 +1222,7 @@ let resolution env full_reified_goal systems_list = (* variables a introduire *) let to_introduce = add_stated_equations env solution_tree in let stated_vars = List.map (fun (v,_,_,_) -> v) to_introduce in - let l_generalize_arg = List.map (fun (_,t,_,_) -> t) to_introduce in + let l_generalize_arg = List.map (fun (_,t,_,_) -> EConstr.of_constr t) to_introduce in let hyp_stated_vars = List.map (fun (_,_,_,id) -> CCEqua id) to_introduce in (* L'environnement de base se construit en deux morceaux : - les variables des équations utiles (et de la conclusion) @@ -1258,6 +1258,7 @@ let resolution env full_reified_goal systems_list = let reified = app coq_interp_sequent [| reified_concl;env_props_reified;env_terms_reified;reified_goal|] in + let reified = EConstr.of_constr reified in let normalize_equation e = let rec loop = function [] -> app (if e.e_negated then coq_p_invert else coq_p_step) @@ -1281,9 +1282,9 @@ let resolution env full_reified_goal systems_list = let decompose_tactic = decompose_tree env context solution_tree in Proofview.V82.of_tactic (Tactics.generalize - (l_generalize_arg @ List.map Term.mkVar (List.tl l_hyps))) >> + (l_generalize_arg @ List.map EConstr.mkVar (List.tl l_hyps))) >> Proofview.V82.of_tactic (Tactics.change_concl reified) >> - Proofview.V82.of_tactic (Tactics.apply (app coq_do_omega [|decompose_tactic; normalization_trace|])) >> + Proofview.V82.of_tactic (Tactics.apply (EConstr.of_constr (app coq_do_omega [|decompose_tactic; normalization_trace|]))) >> show_goal >> Proofview.V82.of_tactic (Tactics.normalise_vm_in_concl) >> (*i Alternatives to the previous line: @@ -1292,7 +1293,7 @@ let resolution env full_reified_goal systems_list = - Skip the conversion check and rely directly on the QED: Tacmach.convert_concl_no_check (Lazy.force coq_True) Term.VMcast >> i*) - Proofview.V82.of_tactic (Tactics.apply (Lazy.force coq_I)) + Proofview.V82.of_tactic (Tactics.apply (EConstr.of_constr (Lazy.force coq_I))) let total_reflexive_omega_tactic gl = Coqlib.check_required_library ["Coq";"romega";"ROmega"]; diff --git a/plugins/rtauto/refl_tauto.ml b/plugins/rtauto/refl_tauto.ml index f88b3a700..981ce2a61 100644 --- a/plugins/rtauto/refl_tauto.ml +++ b/plugins/rtauto/refl_tauto.ml @@ -312,6 +312,7 @@ let rtauto_tac gls= str "Giving proof term to Coq ... ") end in let tac_start_time = System.get_time () in + let term = EConstr.of_constr term in let result= if !check then Proofview.V82.of_tactic (Tactics.exact_check term) gls diff --git a/plugins/ssrmatching/ssrmatching.ml4 b/plugins/ssrmatching/ssrmatching.ml4 index ace557a52..aa91494eb 100644 --- a/plugins/ssrmatching/ssrmatching.ml4 +++ b/plugins/ssrmatching/ssrmatching.ml4 @@ -1392,6 +1392,7 @@ let ssrpatterntac _ist (arg_ist,arg) gl = fill_occ_pattern (Global.env()) sigma0 concl0 pat noindex 1 in let gl, tty = pf_type_of gl (EConstr.of_constr t) in let concl = mkLetIn (Name (id_of_string "selected"), t, tty, concl_x) in + let concl = EConstr.of_constr concl in Proofview.V82.of_tactic (convert_concl concl DEFAULTcast) gl (* Register "ssrpattern" tactic *) diff --git a/pretyping/cases.ml b/pretyping/cases.ml index 57d12a19f..360199fec 100644 --- a/pretyping/cases.ml +++ b/pretyping/cases.ml @@ -297,7 +297,7 @@ let inductive_template evdref env tmloc ind = | LocalAssum (na,ty) -> let ty = EConstr.of_constr ty in let ty' = substl subst ty in - let e = EConstr.of_constr (e_new_evar env evdref ~src:(hole_source n) ty') in + let e = e_new_evar env evdref ~src:(hole_source n) ty' in (e::subst,e::evarl,n+1) | LocalDef (na,b,ty) -> let b = EConstr.of_constr b in @@ -380,7 +380,7 @@ let coerce_to_indtype typing_fun evdref env matx tomatchl = (* Utils *) let mkExistential env ?(src=(Loc.ghost,Evar_kinds.InternalHole)) evdref = - let e, u = e_new_type_evar env evdref univ_flexible_alg ~src:src in EConstr.of_constr e + let e, u = e_new_type_evar env evdref univ_flexible_alg ~src:src in e let evd_comb2 f evdref x y = let (evd',y) = f !evdref x y in @@ -1663,7 +1663,6 @@ let abstract_tycon loc env evdref subst tycon extenv t = 1 (rel_context env) in let ty = EConstr.of_constr ty in let ev' = e_new_evar env evdref ~src ty in - let ev' = EConstr.of_constr ev' in begin match solve_simple_eqn (evar_conv_x full_transparent_state) env !evdref (None,ev,substl inst ev') with | Success evd -> evdref := evd | UnifFailure _ -> assert false @@ -1698,7 +1697,6 @@ let abstract_tycon loc env evdref subst tycon extenv t = let candidates = u :: List.map mkRel vl in let candidates = List.map EConstr.Unsafe.to_constr candidates in let ev = e_new_evar extenv evdref ~src ~filter ~candidates ty in - let ev = EConstr.of_constr ev in lift k ev in aux (0,extenv,subst0) t0 @@ -1712,7 +1710,6 @@ let build_tycon loc env tycon_env s subst tycon extenv evdref t = let n' = Context.Rel.length (rel_context tycon_env) in let impossible_case_type, u = e_new_type_evar (reset_context env) evdref univ_flexible_alg ~src:(loc,Evar_kinds.ImpossibleCase) in - let impossible_case_type = EConstr.of_constr impossible_case_type in (lift (n'-n) impossible_case_type, mkSort u) | Some t -> let t = abstract_tycon loc tycon_env evdref subst tycon extenv t in @@ -2010,7 +2007,7 @@ let prepare_predicate loc typing_fun env sigma tomatchs arsign tycon pred = let Sigma ((t, _), sigma, _) = new_type_evar env sigma univ_flexible_alg ~src:(loc, Evar_kinds.CasesType false) in let sigma = Sigma.to_evar_map sigma in - sigma, EConstr.of_constr t + sigma, t in (* First strategy: we build an "inversion" predicate *) let sigma1,pred1 = build_inversion_problem loc env sigma tomatchs t in diff --git a/pretyping/classops.ml b/pretyping/classops.ml index 9011186a3..23d20dad3 100644 --- a/pretyping/classops.ml +++ b/pretyping/classops.ml @@ -51,6 +51,7 @@ type coe_info_typ = { coe_param : int } let coe_info_typ_equal c1 c2 = + let eq_constr c1 c2 = Termops.eq_constr Evd.empty (EConstr.of_constr c1) (EConstr.of_constr c2) in eq_constr c1.coe_value c2.coe_value && eq_constr c1.coe_type c2.coe_type && c1.coe_local == c2.coe_local && diff --git a/pretyping/coercion.ml b/pretyping/coercion.ml index e7279df7a..d67976232 100644 --- a/pretyping/coercion.ml +++ b/pretyping/coercion.ml @@ -93,7 +93,7 @@ open Program let make_existential loc ?(opaque = not (get_proofs_transparency ())) env evdref c = let src = (loc, Evar_kinds.QuestionMark (Evar_kinds.Define opaque)) in - EConstr.of_constr (Evarutil.e_new_evar env evdref ~src c) + Evarutil.e_new_evar env evdref ~src c let app_opt env evdref f t = EConstr.of_constr (whd_betaiota !evdref (app_opt f t)) diff --git a/pretyping/detyping.ml b/pretyping/detyping.ml index 4756ec30e..ec8945e85 100644 --- a/pretyping/detyping.ml +++ b/pretyping/detyping.ml @@ -294,7 +294,7 @@ and align_tree nal isgoal (e,c as rhs) sigma = match nal with | na::nal -> match kind_of_term c with | Case (ci,p,c,cl) when - eq_constr c (mkRel (List.index Name.equal na (fst (snd e)))) + eq_constr sigma (EConstr.of_constr c) (EConstr.mkRel (List.index Name.equal na (fst (snd e)))) && not (Int.equal (Array.length cl) 0) && (* don't contract if p dependent *) computable p (List.length ci.ci_pp_info.ind_tags) (* FIXME: can do better *) -> diff --git a/pretyping/evarconv.ml b/pretyping/evarconv.ml index ee6355736..a968af7ff 100644 --- a/pretyping/evarconv.ml +++ b/pretyping/evarconv.ml @@ -890,7 +890,7 @@ and conv_record trs env evd (ctx,(h,h2),c,bs,(params,params1),(us,us2),(sk1,sk2) let i = Sigma.Unsafe.of_evar_map i in let Sigma (ev, i', _) = Evarutil.new_evar env i ~src:dloc (substl ks b) in let i' = Sigma.to_evar_map i' in - (i', EConstr.of_constr ev :: ks, m - 1,test)) + (i', ev :: ks, m - 1,test)) (evd,[],List.length bs,fun i -> Success i) bs in let app = mkApp (c, Array.rev_of_list ks) in @@ -1066,13 +1066,13 @@ let second_order_matching ts env_rhs evd (evk,args) argoccs rhs = | Some _ -> error "Selection of specific occurrences not supported" | None -> let evty = set_holes evdref cty subst in - let instance = List.map EConstr.Unsafe.to_constr (Filter.filter_list filter instance) in + let instance = Filter.filter_list filter instance in let evd = Sigma.Unsafe.of_evar_map !evdref in let Sigma (ev, evd, _) = new_evar_instance sign evd evty ~filter instance in let evd = Sigma.to_evar_map evd in evdref := evd; - evsref := (fst (destEvar !evdref (EConstr.of_constr ev)),evty)::!evsref; - EConstr.of_constr ev in + evsref := (fst (destEvar !evdref ev),evty)::!evsref; + ev in set_holes evdref (apply_on_subterm env_rhs evdref set_var c rhs) subst | [] -> rhs in diff --git a/pretyping/evardefine.ml b/pretyping/evardefine.ml index fa3b9ca0b..3babc48a7 100644 --- a/pretyping/evardefine.ml +++ b/pretyping/evardefine.ml @@ -22,6 +22,11 @@ open Sigma.Notations module RelDecl = Context.Rel.Declaration +let nlocal_assum (na, t) = + let open Context.Named.Declaration in + let inj = EConstr.Unsafe.to_constr in + LocalAssum (na, inj t) + 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 @@ -88,7 +93,7 @@ let define_pure_evar_as_product evd evk = (Sigma.to_evar_map evd1, e) in let evd2,rng = - let newenv = push_named (LocalAssum (id, dom)) evenv in + let newenv = push_named (nlocal_assum (id, dom)) evenv in let src = evar_source evk evd1 in let filter = Filter.extend 1 (evar_filter evi) in if is_prop_sort s then @@ -105,8 +110,7 @@ let define_pure_evar_as_product evd evk = let evd3 = Evd.set_leq_sort evenv evd3 (Type prods) s in evd3, rng in - let rng = EConstr.of_constr rng in - let prod = mkProd (Name id, EConstr.of_constr dom, subst_var id rng) in + let prod = mkProd (Name id, dom, subst_var id rng) in let evd3 = Evd.define evk (EConstr.Unsafe.to_constr prod) evd2 in evd3,prod @@ -140,14 +144,13 @@ let define_pure_evar_as_lambda env evd evk = | Evar ev' -> let evd,typ = define_evar_as_product evd ev' in evd,destProd evd typ | _ -> error_not_product env evd typ in let avoid = ids_of_named_context (evar_context evi) in - let dom = EConstr.Unsafe.to_constr dom in let id = - next_name_away_with_default_using_types "x" na avoid (Reductionops.whd_evar evd dom) in - let newenv = push_named (LocalAssum (id, dom)) evenv in + next_name_away_with_default_using_types "x" na avoid (Reductionops.whd_evar evd (EConstr.Unsafe.to_constr dom)) in + let newenv = push_named (nlocal_assum (id, dom)) evenv in let filter = Filter.extend 1 (evar_filter evi) in let src = evar_source evk evd1 in let evd2,body = new_evar_unsafe newenv evd1 ~src (subst1 (mkVar id) rng) ~filter in - let lam = mkLambda (Name id, EConstr.of_constr dom, subst_var id (EConstr.of_constr body)) in + let lam = mkLambda (Name id, dom, subst_var id body) in Evd.define evk (EConstr.Unsafe.to_constr lam) evd2, lam let define_evar_as_lambda env evd (evk,args) = diff --git a/pretyping/evarsolve.ml b/pretyping/evarsolve.ml index b7db51d5c..4ce8a44ad 100644 --- a/pretyping/evarsolve.ml +++ b/pretyping/evarsolve.ml @@ -518,15 +518,15 @@ let is_unification_pattern (env,nb) evd f l t = let solve_pattern_eqn env sigma l c = let c' = List.fold_right (fun a c -> let c' = subst_term sigma (lift 1 a) (lift 1 c) in + let c' = EConstr.of_constr c' in match EConstr.kind sigma a with (* Rem: if [a] links to a let-in, do as if it were an assumption *) | Rel n -> let open Context.Rel.Declaration in let d = map_constr (CVars.lift n) (lookup_rel n env) in - let c' = EConstr.of_constr c' in mkLambda_or_LetIn d c' | Var id -> - let d = lookup_named id env in EConstr.of_constr (mkNamedLambda_or_LetIn d c') + let d = lookup_named id env in mkNamedLambda_or_LetIn d c' | _ -> assert false) l c in (* Warning: we may miss some opportunity to eta-reduce more since c' @@ -592,10 +592,9 @@ let make_projectable_subst aliases sigma evi args = let define_evar_from_virtual_equation define_fun env evd src t_in_env ty_t_in_sign sign filter inst_in_env = let evd = Sigma.Unsafe.of_evar_map evd in - let Sigma (evar_in_env, evd, _) = new_evar_instance sign evd (EConstr.of_constr ty_t_in_sign) ~filter ~src (List.map EConstr.Unsafe.to_constr inst_in_env) in + let Sigma (evar_in_env, evd, _) = new_evar_instance sign evd (EConstr.of_constr ty_t_in_sign) ~filter ~src inst_in_env in let evd = Sigma.to_evar_map evd in let t_in_env = EConstr.of_constr (whd_evar evd (EConstr.Unsafe.to_constr t_in_env)) in - let evar_in_env = EConstr.of_constr evar_in_env in let (evk, _) = destEvar evd evar_in_env in let evd = define_fun env evd None (EConstr.destEvar evd evar_in_env) t_in_env in let ctxt = named_context_of_val sign in @@ -669,10 +668,10 @@ let materialize_evar define_fun env evd k (evk1,args1) ty_in_env = let evd = Sigma.Unsafe.of_evar_map evd in let ev2ty_in_sign = EConstr.of_constr ev2ty_in_sign in let Sigma (ev2_in_sign, evd, _) = - new_evar_instance sign2 evd ev2ty_in_sign ~filter:filter2 ~src (List.map EConstr.Unsafe.to_constr inst2_in_sign) in + new_evar_instance sign2 evd ev2ty_in_sign ~filter:filter2 ~src inst2_in_sign in let evd = Sigma.to_evar_map evd in - let ev2_in_env = (fst (destEvar evd (EConstr.of_constr ev2_in_sign)), Array.of_list inst2_in_env) in - (evd, EConstr.of_constr ev2_in_sign, ev2_in_env) + let ev2_in_env = (fst (destEvar evd ev2_in_sign), Array.of_list inst2_in_env) in + (evd, ev2_in_sign, ev2_in_env) let restrict_upon_filter evd evk p args = let oldfullfilter = evar_filter (Evd.find_undefined evd evk) in diff --git a/pretyping/inductiveops.ml b/pretyping/inductiveops.ml index e30ba21fd..98b267cfd 100644 --- a/pretyping/inductiveops.ml +++ b/pretyping/inductiveops.ml @@ -556,6 +556,7 @@ let set_pattern_names env ind brv = let type_case_branches_with_names env sigma indspec p c = let (ind,args) = indspec in + let args = List.map EConstr.Unsafe.to_constr args in let (mib,mip as specif) = Inductive.lookup_mind_specif env (fst ind) in let nparams = mib.mind_nparams in let (params,realargs) = List.chop nparams args in diff --git a/pretyping/inductiveops.mli b/pretyping/inductiveops.mli index cf5523a50..7af9731f9 100644 --- a/pretyping/inductiveops.mli +++ b/pretyping/inductiveops.mli @@ -175,7 +175,7 @@ val arity_of_case_predicate : env -> inductive_family -> bool -> sorts -> types val type_case_branches_with_names : - env -> evar_map -> pinductive * constr list -> constr -> constr -> types array * types + env -> evar_map -> pinductive * EConstr.constr list -> constr -> constr -> types array * types (** Annotation for cases *) val make_case_info : env -> inductive -> case_style -> case_info diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml index 49a0bccee..7586b54ba 100644 --- a/pretyping/pretyping.ml +++ b/pretyping/pretyping.ml @@ -116,7 +116,7 @@ let lookup_named id env = lookup_named id env.env let e_new_evar env evdref ?src ?naming typ = let subst2 subst vsubst c = csubst_subst subst (replace_vars vsubst c) in let open Context.Named.Declaration in - let inst_vars = List.map (get_id %> Constr.mkVar) (named_context env.env) in + let inst_vars = List.map (get_id %> EConstr.mkVar) (named_context env.env) in let inst_rels = List.rev (rel_list 0 (nb_rel env.env)) in let (subst, vsubst, _, nc) = Lazy.force env.extra in let typ' = subst2 subst vsubst typ in @@ -125,7 +125,7 @@ let e_new_evar env evdref ?src ?naming typ = let sigma = Sigma.Unsafe.of_evar_map !evdref in let Sigma (e, sigma, _) = new_evar_instance sign sigma typ' ?src ?naming instance in evdref := Sigma.to_evar_map sigma; - EConstr.of_constr e + e let push_rec_types (lna,typarray,_) env = let ctxt = Array.map2_i (fun i na t -> local_assum (na, lift i t)) lna typarray in @@ -546,7 +546,7 @@ let new_type_evar env evdref loc = univ_flexible_alg ~src:(loc,Evar_kinds.InternalHole) in evdref := Sigma.to_evar_map sigma; - EConstr.of_constr e + e let (f_genarg_interp, genarg_interp_hook) = Hook.make () diff --git a/pretyping/tacred.ml b/pretyping/tacred.ml index 24d4af89a..1ec8deb1b 100644 --- a/pretyping/tacred.ml +++ b/pretyping/tacred.ml @@ -1196,7 +1196,7 @@ let reduce_to_ind_gen allow_product env sigma t = let t = hnf_constr env sigma t in let t = EConstr.of_constr t in match EConstr.kind sigma (EConstr.of_constr (fst (decompose_app_vect sigma t))) with - | Ind ind-> (check_privacy env ind, EConstr.Unsafe.to_constr (it_mkProd_or_LetIn t l)) + | Ind ind-> (check_privacy env ind, it_mkProd_or_LetIn t l) | Prod (n,ty,t') -> let open Context.Rel.Declaration in if allow_product then @@ -1209,7 +1209,7 @@ let reduce_to_ind_gen allow_product env sigma t = let t' = whd_all env sigma t in let t' = EConstr.of_constr t' in match EConstr.kind sigma (EConstr.of_constr (fst (decompose_app_vect sigma t'))) with - | Ind ind-> (check_privacy env ind, EConstr.Unsafe.to_constr (it_mkProd_or_LetIn t' l)) + | Ind ind-> (check_privacy env ind, it_mkProd_or_LetIn t' l) | _ -> user_err (str"Not an inductive product.") in elimrec env t [] @@ -1219,7 +1219,7 @@ let reduce_to_atomic_ind env sigma c = reduce_to_ind_gen false env sigma c let find_hnf_rectype env sigma t = let ind,t = reduce_to_atomic_ind env sigma t in - ind, snd (Term.decompose_app t) + ind, snd (decompose_app sigma t) (* Reduce the weak-head redex [beta,iota/fix/cofix[all],cast,zeta,simpl/delta] or raise [NotStepReducible] if not a weak-head redex *) @@ -1295,7 +1295,7 @@ let reduce_to_ref_gen allow_product env sigma ref t = elimrec env t' l with NotStepReducible -> error_cannot_recognize ref in - EConstr.Unsafe.to_constr (elimrec env t []) + elimrec env t [] let reduce_to_quantified_ref = reduce_to_ref_gen true let reduce_to_atomic_ref = reduce_to_ref_gen false diff --git a/pretyping/tacred.mli b/pretyping/tacred.mli index 3587ae281..15b4e990d 100644 --- a/pretyping/tacred.mli +++ b/pretyping/tacred.mli @@ -75,23 +75,23 @@ val cbv_norm_flags : CClosure.RedFlags.reds -> reduction_function (** [reduce_to_atomic_ind env sigma t] puts [t] in the form [t'=(I args)] with [I] an inductive definition; returns [I] and [t'] or fails with a user error *) -val reduce_to_atomic_ind : env -> evar_map -> EConstr.types -> pinductive * types +val reduce_to_atomic_ind : env -> evar_map -> EConstr.types -> pinductive * EConstr.types (** [reduce_to_quantified_ind env sigma t] puts [t] in the form [t'=(x1:A1)..(xn:An)(I args)] with [I] an inductive definition; returns [I] and [t'] or fails with a user error *) -val reduce_to_quantified_ind : env -> evar_map -> EConstr.types -> pinductive * types +val reduce_to_quantified_ind : env -> evar_map -> EConstr.types -> pinductive * EConstr.types (** [reduce_to_quantified_ref env sigma ref t] try to put [t] in the form [t'=(x1:A1)..(xn:An)(ref args)] and fails with user error if not possible *) val reduce_to_quantified_ref : - env -> evar_map -> global_reference -> EConstr.types -> types + env -> evar_map -> global_reference -> EConstr.types -> EConstr.types val reduce_to_atomic_ref : - env -> evar_map -> global_reference -> EConstr.types -> types + env -> evar_map -> global_reference -> EConstr.types -> EConstr.types val find_hnf_rectype : - env -> evar_map -> EConstr.types -> pinductive * constr list + env -> evar_map -> EConstr.types -> pinductive * EConstr.constr list val contextually : bool -> occurrences * constr_pattern -> (patvar_map -> reduction_function) -> reduction_function diff --git a/pretyping/unification.ml b/pretyping/unification.ml index 2b2069ec4..bc59a4108 100644 --- a/pretyping/unification.ml +++ b/pretyping/unification.ml @@ -133,14 +133,14 @@ let abstract_list_all_with_dependencies env evd typ c l = let evd = Sigma.Unsafe.of_evar_map evd in let Sigma (ev, evd, _) = new_evar env evd typ in let evd = Sigma.to_evar_map evd in - let evd,ev' = evar_absorb_arguments env evd (destEvar evd (EConstr.of_constr ev)) l in + let evd,ev' = evar_absorb_arguments env evd (destEvar evd ev) l in let n = List.length l in let argoccs = set_occurrences_of_last_arg (Array.sub (snd ev') 0 n) in let evd,b = Evarconv.second_order_matching empty_transparent_state env evd ev' argoccs c in if b then - let p = nf_evar evd ev in + let p = nf_evar evd (EConstr.Unsafe.to_constr ev) in evd, p else error_cannot_find_well_typed_abstraction env evd c l None @@ -184,8 +184,8 @@ let pose_all_metas_as_evars env evd t = let ty = if Evd.Metaset.is_empty mvs then ty else aux ty in let src = Evd.evar_source_of_meta mv !evdref in let ev = Evarutil.e_new_evar env evdref ~src ty in - evdref := meta_assign mv (ev,(Conv,TypeNotProcessed)) !evdref; - EConstr.of_constr ev) + evdref := meta_assign mv (EConstr.Unsafe.to_constr ev,(Conv,TypeNotProcessed)) !evdref; + ev) | _ -> EConstr.map !evdref aux t in let c = aux t in @@ -1236,7 +1236,6 @@ let applyHead env (type r) (evd : r Sigma.t) n c = match EConstr.kind sigma (EConstr.of_constr (whd_all env sigma cty)) with | Prod (_,c1,c2) -> let Sigma (evar, evd', q) = Evarutil.new_evar env evd ~src:(Loc.ghost,Evar_kinds.GoalEvar) c1 in - let evar = EConstr.of_constr evar in apprec (n-1) (mkApp(c,[|evar|])) (subst1 evar c2) (p +> q) evd' | _ -> error "Apply_Head_Then" in diff --git a/proofs/clenv.ml b/proofs/clenv.ml index fd88e3c51..514fc27e8 100644 --- a/proofs/clenv.ml +++ b/proofs/clenv.ml @@ -337,7 +337,6 @@ let clenv_pose_metas_as_evars clenv dep_mvs = let evd = Sigma.Unsafe.of_evar_map clenv.evd in let Sigma (evar, evd, _) = new_evar (cl_env clenv) evd ~src ty in let evd = Sigma.to_evar_map evd in - let evar = EConstr.of_constr evar in let clenv = clenv_assign mv evar {clenv with evd=evd} in fold clenv mvs in fold clenv dep_mvs @@ -619,7 +618,6 @@ let make_evar_clause env sigma ?len t = let sigma = Sigma.Unsafe.of_evar_map sigma in let Sigma (ev, sigma, _) = new_evar ~store env sigma t1 in let sigma = Sigma.to_evar_map sigma in - let ev = EConstr.of_constr ev in let dep = not (noccurn sigma 1 t2) in let hole = { hole_evar = ev; diff --git a/proofs/tacmach.mli b/proofs/tacmach.mli index 07d02212c..cfbfe12b1 100644 --- a/proofs/tacmach.mli +++ b/proofs/tacmach.mli @@ -67,8 +67,8 @@ val pf_whd_all : goal sigma -> EConstr.constr -> constr val pf_hnf_constr : goal sigma -> EConstr.constr -> constr val pf_nf : goal sigma -> EConstr.constr -> constr val pf_nf_betaiota : goal sigma -> EConstr.constr -> constr -val pf_reduce_to_quantified_ind : goal sigma -> EConstr.types -> pinductive * types -val pf_reduce_to_atomic_ind : goal sigma -> EConstr.types -> pinductive * types +val pf_reduce_to_quantified_ind : goal sigma -> EConstr.types -> pinductive * EConstr.types +val pf_reduce_to_atomic_ind : goal sigma -> EConstr.types -> pinductive * EConstr.types val pf_compute : goal sigma -> EConstr.constr -> constr val pf_unfoldn : (occurrences * evaluable_global_reference) list -> goal sigma -> EConstr.constr -> constr @@ -121,7 +121,7 @@ module New : sig val pf_last_hyp : ([ `NF ], 'r) Proofview.Goal.t -> Context.Named.Declaration.t val pf_nf_concl : ([ `LZ ], 'r) Proofview.Goal.t -> types - val pf_reduce_to_quantified_ind : ('a, 'r) Proofview.Goal.t -> EConstr.types -> pinductive * types + val pf_reduce_to_quantified_ind : ('a, 'r) Proofview.Goal.t -> EConstr.types -> pinductive * EConstr.types val pf_hnf_constr : ('a, 'r) Proofview.Goal.t -> EConstr.constr -> types val pf_hnf_type_of : ('a, 'r) Proofview.Goal.t -> EConstr.constr -> types diff --git a/stm/lemmas.ml b/stm/lemmas.ml index 9896d5a93..04f888a84 100644 --- a/stm/lemmas.ml +++ b/stm/lemmas.ml @@ -393,7 +393,7 @@ let start_proof_univs id ?pl kind sigma ?terminator ?sign c ?init_tac ?(compute_ let rec_tac_initializer finite guard thms snl = if finite then - match List.map (fun ((id,_),(t,_)) -> (id,t)) thms with + match List.map (fun ((id,_),(t,_)) -> (id,EConstr.of_constr t)) thms with | (id,_)::l -> Tactics.mutual_cofix id l 0 | _ -> assert false else @@ -401,7 +401,7 @@ let rec_tac_initializer finite guard thms snl = let nl = match snl with | None -> List.map succ (List.map List.last guard) | Some nl -> nl - in match List.map2 (fun ((id,_),(t,_)) n -> (id,n,t)) thms nl with + in match List.map2 (fun ((id,_),(t,_)) n -> (id,n, EConstr.of_constr t)) thms nl with | (id,n,_)::l -> Tactics.mutual_fix id n l 0 | _ -> assert false diff --git a/stm/stm.ml b/stm/stm.ml index 6012e3d2d..d60412c0c 100644 --- a/stm/stm.ml +++ b/stm/stm.ml @@ -1792,7 +1792,7 @@ end = struct (* {{{ *) str"uc=" ++ Evd.pr_evar_universe_context uc))); (if abstract then Tactics.tclABSTRACT None else (fun x -> x)) (V82.tactic (Refiner.tclPUSHEVARUNIVCONTEXT uc) <*> - Tactics.exact_no_check pt) + Tactics.exact_no_check (EConstr.of_constr pt)) with TacTask.NoProgress -> if solve then Tacticals.New.tclSOLVE [] else tclUNIT () }) diff --git a/tactics/auto.ml b/tactics/auto.ml index 2b654f563..41b56bd3d 100644 --- a/tactics/auto.ml +++ b/tactics/auto.ml @@ -115,6 +115,7 @@ let unify_resolve_gen poly = function let exact poly (c,clenv) = Proofview.Goal.enter { enter = begin fun gl -> let clenv', c = connect_hint_clenv poly c clenv gl in + let c = EConstr.of_constr c in Tacticals.New.tclTHEN (Proofview.Unsafe.tclEVARUNIVCONTEXT (Evd.evar_universe_context clenv'.evd)) (exact_check c) diff --git a/tactics/class_tactics.ml b/tactics/class_tactics.ml index a8768b6ed..7d8fc79f4 100644 --- a/tactics/class_tactics.ml +++ b/tactics/class_tactics.ml @@ -226,7 +226,8 @@ let e_give_exact flags poly (c,clenv) gl = c, {gl with sigma = evd} else c, gl in - let t1 = pf_unsafe_type_of gl (EConstr.of_constr c) in + let c = EConstr.of_constr c in + let t1 = pf_unsafe_type_of gl c in let t1 = EConstr.of_constr t1 in Proofview.V82.of_tactic (Clenvtac.unify ~flags t1 <*> exact_no_check c) gl @@ -1483,7 +1484,7 @@ let resolve_one_typeclass env ?(sigma=Evd.empty) gl unique = let evd = sig_sig gls' in let t = EConstr.Unsafe.to_constr t in let t' = let (ev, inst) = destEvar t in - mkEvar (ev, Array.of_list subst) + mkEvar (ev, Array.map_of_list EConstr.Unsafe.to_constr subst) in let term = Evarutil.nf_evar evd t' in evd, term @@ -1506,6 +1507,7 @@ let rec head_of_constr sigma t = let head_of_constr h c = Proofview.tclEVARMAP >>= fun sigma -> let c = head_of_constr sigma c in + let c = EConstr.of_constr c in letin_tac None (Name h) c None Locusops.allHyps let not_evar c = diff --git a/tactics/contradiction.ml b/tactics/contradiction.ml index a8be704b2..a92b14dbe 100644 --- a/tactics/contradiction.ml +++ b/tactics/contradiction.ml @@ -33,8 +33,8 @@ let absurd c = let t = EConstr.Unsafe.to_constr j.Environ.utj_val in let tac = Tacticals.New.tclTHENLIST [ - elim_type (build_coq_False ()); - Simple.apply (mk_absurd_proof t) + elim_type (EConstr.of_constr (build_coq_False ())); + Simple.apply (EConstr.of_constr (mk_absurd_proof t)) ] in Sigma.Unsafe.of_pair (tac, sigma) end } @@ -67,7 +67,7 @@ let contradiction_context = let typ = nf_evar sigma (NamedDecl.get_type d) in let typ = whd_all env sigma (EConstr.of_constr typ) in if is_empty_type sigma (EConstr.of_constr typ) then - simplest_elim (mkVar id) + simplest_elim (EConstr.mkVar id) else match kind_of_term typ with | Prod (na,t,u) when is_empty_type sigma (EConstr.of_constr u) -> let is_unit_or_eq = @@ -82,14 +82,14 @@ let contradiction_context = let params = Util.List.firstn nparams args in let p = applist ((mkConstructUi (indu,1)), params) in (* Checking on the fly that it type-checks *) - simplest_elim (mkApp (mkVar id,[|p|])) + simplest_elim (EConstr.mkApp (EConstr.mkVar id,[|EConstr.of_constr p|])) | None -> Tacticals.New.tclZEROMSG (Pp.str"Not a negated unit type.")) (Proofview.tclORELSE (Proofview.Goal.enter { enter = begin fun gl -> let is_conv_leq = Tacmach.New.pf_apply is_conv_leq gl in filter_hyp (fun typ -> is_conv_leq (EConstr.of_constr typ) (EConstr.of_constr t)) - (fun id' -> simplest_elim (mkApp (mkVar id,[|mkVar id'|]))) + (fun id' -> simplest_elim (EConstr.mkApp (EConstr.mkVar id,[|EConstr.mkVar id'|]))) end }) begin function (e, info) -> match e with | Not_found -> seek_neg rest @@ -113,7 +113,7 @@ let contradiction_term (c,lbind as cl) = let sigma = Tacmach.New.project gl in let env = Proofview.Goal.env gl in let type_of = Tacmach.New.pf_unsafe_type_of gl in - let typ = type_of (EConstr.of_constr c) in + let typ = type_of c in let _, ccl = splay_prod env sigma (EConstr.of_constr typ) in if is_empty_type sigma (EConstr.of_constr ccl) then Tacticals.New.tclTHEN @@ -124,7 +124,7 @@ let contradiction_term (c,lbind as cl) = begin if lbind = NoBindings then filter_hyp (fun c -> is_negation_of env sigma typ (EConstr.of_constr c)) - (fun id -> simplest_elim (mkApp (mkVar id,[|c|]))) + (fun id -> simplest_elim (EConstr.mkApp (EConstr.mkVar id,[|c|]))) else Proofview.tclZERO Not_found end diff --git a/tactics/contradiction.mli b/tactics/contradiction.mli index b876aee90..5cc4b2e01 100644 --- a/tactics/contradiction.mli +++ b/tactics/contradiction.mli @@ -10,4 +10,4 @@ open Term open Misctypes val absurd : constr -> unit Proofview.tactic -val contradiction : constr with_bindings option -> unit Proofview.tactic +val contradiction : EConstr.constr with_bindings option -> unit Proofview.tactic diff --git a/tactics/eauto.ml b/tactics/eauto.ml index 7b07c9309..24e4de750 100644 --- a/tactics/eauto.ml +++ b/tactics/eauto.ml @@ -29,8 +29,9 @@ open Proofview.Notations let eauto_unif_flags = auto_flags_of_state full_transparent_state let e_give_exact ?(flags=eauto_unif_flags) c = + let c = EConstr.of_constr c in Proofview.Goal.enter { enter = begin fun gl -> - let t1 = Tacmach.New.pf_unsafe_type_of gl (EConstr.of_constr c) in + let t1 = Tacmach.New.pf_unsafe_type_of gl c in let t1 = EConstr.of_constr t1 in let t2 = Tacmach.New.pf_concl (Proofview.Goal.assume gl) in let sigma = Tacmach.New.project gl in @@ -77,7 +78,7 @@ let apply_tac_list tac glls = let one_step l gl = [Proofview.V82.of_tactic Tactics.intro] - @ (List.map (fun c -> Proofview.V82.of_tactic (Tactics.Simple.eapply c)) (List.map mkVar (pf_ids_of_hyps gl))) + @ (List.map (fun c -> Proofview.V82.of_tactic (Tactics.Simple.eapply c)) (List.map EConstr.mkVar (pf_ids_of_hyps gl))) @ (List.map (fun c -> Proofview.V82.of_tactic (Tactics.Simple.eapply c)) l) @ (List.map (fun c -> Proofview.V82.of_tactic (assumption c)) (pf_ids_of_hyps gl)) @@ -94,8 +95,9 @@ let prolog_tac l n = Proofview.V82.tactic begin fun gl -> let map c = let (c, sigma) = Tactics.run_delayed (pf_env gl) (project gl) c in + let c = EConstr.Unsafe.to_constr c in let c = pf_apply (prepare_hint false (false,true)) gl (sigma, c) in - out_term c + EConstr.of_constr (out_term c) in let l = List.map map l in try (prolog l n gl) @@ -114,6 +116,7 @@ let priority l = List.map snd (List.filter (fun (pr,_) -> Int.equal pr 0) l) let unify_e_resolve poly flags (c,clenv) = Proofview.Goal.nf_enter { enter = begin fun gl -> let clenv', c = connect_hint_clenv poly c clenv gl in + let c = EConstr.of_constr c in Proofview.V82.tactic (fun gls -> let clenv' = clenv_unique_resolver ~flags clenv' gls in @@ -515,6 +518,7 @@ let autounfold_one db cl = let did, c' = unfold_head env st (match cl with Some (id, _) -> Tacmach.New.pf_get_hyp_typ id gl | None -> concl) in + let c' = EConstr.of_constr c' in if did then match cl with | Some hyp -> change_in_hyp None (make_change_arg c') hyp diff --git a/tactics/elim.ml b/tactics/elim.ml index d00e504ff..e641f970a 100644 --- a/tactics/elim.ml +++ b/tactics/elim.ml @@ -77,10 +77,12 @@ let tmphyp_name = Id.of_string "_TmpHyp" let up_to_delta = ref false (* true *) let general_decompose recognizer c = + let c = EConstr.of_constr c in Proofview.Goal.enter { enter = begin fun gl -> let type_of = pf_unsafe_type_of gl in let sigma = project gl in - let typc = type_of (EConstr.of_constr c) in + let typc = type_of c in + let typc = EConstr.of_constr typc in tclTHENS (cut typc) [ tclTHEN (intro_using tmphyp_name) (onLastHypId diff --git a/tactics/eqdecide.ml b/tactics/eqdecide.ml index ed81d748a..eb75cbf7d 100644 --- a/tactics/eqdecide.ml +++ b/tactics/eqdecide.ml @@ -25,6 +25,7 @@ open Misctypes open Tactypes open Hipattern open Pretyping +open Proofview.Notations open Tacmach.New open Coqlib @@ -50,7 +51,10 @@ open Coqlib Eduardo Gimenez (30/3/98). *) -let clear_last = (onLastHyp (fun c -> (clear [destVar c]))) +let clear_last = + let open EConstr in + Proofview.tclEVARMAP >>= fun sigma -> + (onLastHyp (fun c -> (clear [destVar sigma c]))) let choose_eq eqonleft = if eqonleft then @@ -66,14 +70,14 @@ let choose_noteq eqonleft = let mkBranches c1 c2 = tclTHENLIST [generalize [c2]; - Simple.elim c1; + Simple.elim (EConstr.of_constr c1); intros; onLastHyp Simple.case; clear_last; intros] let discrHyp id = - let c = { delayed = fun env sigma -> Sigma.here (Term.mkVar id, NoBindings) sigma } in + let c = { delayed = fun env sigma -> Sigma.here (EConstr.mkVar id, NoBindings) sigma } in let tac c = Equality.discr_tac false (Some (None, ElimOnConstr c)) in Tacticals.New.tclDELAYEDWITHHOLES false c tac @@ -121,7 +125,7 @@ let eqCase tac = tclTHEN intro (onLastHypId tac) let injHyp id = - let c = { delayed = fun env sigma -> Sigma.here (Term.mkVar id, NoBindings) sigma } in + let c = { delayed = fun env sigma -> Sigma.here (EConstr.mkVar id, NoBindings) sigma } in let tac c = Equality.injClause None false (Some (None, ElimOnConstr c)) in Tacticals.New.tclDELAYEDWITHHOLES false c tac @@ -133,7 +137,7 @@ let diseqCase hyps eqonleft = (tclTHEN (rewrite_and_clear (List.rev hyps)) (tclTHEN (red_in_concl) (tclTHEN (intro_using absurd) - (tclTHEN (Simple.apply (mkVar diseq)) + (tclTHEN (Simple.apply (EConstr.mkVar diseq)) (tclTHEN (injHyp absurd) (full_trivial [])))))))) @@ -158,6 +162,7 @@ let rec solveArg hyps eqonleft op largs rargs = match largs, rargs with Proofview.Goal.enter { enter = begin fun gl -> let rectype = pf_unsafe_type_of gl (EConstr.of_constr a1) in let decide = mkDecideEqGoal eqonleft op rectype a1 a2 in + let decide = EConstr.of_constr decide in let tac hyp = solveArg (hyp :: hyps) eqonleft op largs rargs in let subtacs = if eqonleft then [eqCase tac;diseqCase hyps eqonleft;default_auto] @@ -207,7 +212,7 @@ let decideGralEquality = | _ -> tclZEROMSG (Pp.str"This decision procedure only works for inductive objects.") end >>= fun rectype -> (tclTHEN - (mkBranches c1 c2) + (mkBranches c1 (EConstr.of_constr c2)) (tclORELSE (solveNoteqBranch eqonleft) (solveEqBranch rectype))) end } end @@ -222,6 +227,7 @@ let decideEqualityGoal = tclTHEN intros decideGralEquality let decideEquality rectype = Proofview.Goal.enter { enter = begin fun gl -> let decide = mkGenDecideEqGoal rectype gl in + let decide = EConstr.of_constr decide in (tclTHENS (cut decide) [default_auto;decideEqualityGoal]) end } @@ -232,6 +238,7 @@ let compare c1 c2 = Proofview.Goal.enter { enter = begin fun gl -> let rectype = pf_unsafe_type_of gl (EConstr.of_constr c1) in let decide = mkDecideEqGoal true (build_coq_sumbool ()) rectype c1 c2 in + let decide = EConstr.of_constr decide in (tclTHENS (cut decide) [(tclTHEN intro (tclTHEN (onLastHyp simplest_case) clear_last)); diff --git a/tactics/eqschemes.ml b/tactics/eqschemes.ml index 92480e253..57bac25b9 100644 --- a/tactics/eqschemes.ml +++ b/tactics/eqschemes.ml @@ -120,7 +120,7 @@ let get_sym_eq_data env (ind,u) = let paramsctxt = Vars.subst_instance_context u mib.mind_params_ctxt in let paramsctxt1,_ = List.chop (mib.mind_nparams-mip.mind_nrealargs) paramsctxt in - if not (List.equal eq_constr params2 constrargs) then + if not (List.equal Term.eq_constr params2 constrargs) then error "Constructors arguments must repeat the parameters."; (* nrealargs_ctxt and nrealargs are the same here *) (specif,mip.mind_nrealargs,realsign,paramsctxt,paramsctxt1) diff --git a/tactics/equality.ml b/tactics/equality.ml index e1a8d2bdb..80f83f19b 100644 --- a/tactics/equality.ml +++ b/tactics/equality.ml @@ -186,8 +186,8 @@ let instantiate_lemma_all frzevars gl c ty l l2r concl = let instantiate_lemma gl c ty l l2r concl = let c = EConstr.of_constr c in let sigma, ct = pf_type_of gl c in - let t = try snd (reduce_to_quantified_ind (pf_env gl) sigma (EConstr.of_constr ct)) with UserError _ -> ct in - let t = EConstr.of_constr t in + let ct = EConstr.of_constr ct in + let t = try snd (reduce_to_quantified_ind (pf_env gl) sigma ct) with UserError _ -> ct in let l = Miscops.map_bindings EConstr.of_constr l in let eqclause = Clenv.make_clenv_binding (pf_env gl) sigma (c,t) l in [eqclause] @@ -413,6 +413,7 @@ let leibniz_rewrite_ebindings_clause cls lft2rgt tac c t l with_evars frzevars d let type_of_cls = type_of_clause cls gl in let dep = dep_proof_ok && dep_fun evd (EConstr.of_constr c) (EConstr.of_constr type_of_cls) in let Sigma ((elim, effs), sigma, p) = find_elim hdcncl lft2rgt dep cls (Some t) gl in + let elim = EConstr.of_constr elim in let tac = Proofview.tclEFFECTS effs <*> general_elim_clause with_evars frzevars tac cls c t l @@ -562,6 +563,7 @@ let general_multi_rewrite with_evars l cl tac = let sigma = Tacmach.New.project gl in let env = Proofview.Goal.env gl in let (c, sigma) = run_delayed env sigma f in + let c = Miscops.map_with_bindings EConstr.Unsafe.to_constr c in tclWITHHOLES with_evars (general_rewrite_clause l2r with_evars ?tac c cl) sigma end } @@ -646,6 +648,8 @@ let replace_using_leibniz clause c1 c2 l2r unsafe try_prove_eq_opt = Tacticals.New.pf_constr_of_global sym (fun sym -> Tacticals.New.pf_constr_of_global e (fun e -> let eq = applist (e, [t1;c1;c2]) in + let sym = EConstr.of_constr sym in + let eq = EConstr.of_constr eq in tclTHENLAST (replace_core clause l2r eq) (tclFIRST @@ -948,7 +952,7 @@ let gen_absurdity id = let hyp_typ = EConstr.of_constr hyp_typ in if is_empty_type sigma hyp_typ then - simplest_elim (mkVar id) + simplest_elim (EConstr.mkVar id) else tclZEROMSG (str "Not the negation of an equality.") end } @@ -996,6 +1000,7 @@ let discr_positions env sigma (lbeq,eqn,(t,t1,t2)) eq_clause cpath dirn = let t = EConstr.Unsafe.to_constr t in let t1 = EConstr.Unsafe.to_constr t1 in let t2 = EConstr.Unsafe.to_constr t2 in + let eqn = EConstr.Unsafe.to_constr eqn in let e = next_ident_away eq_baseid (ids_of_context env) in let e_env = push_named (Context.Named.Declaration.LocalAssum (e,t)) env in let discriminator = @@ -1004,6 +1009,7 @@ let discr_positions env sigma (lbeq,eqn,(t,t1,t2)) eq_clause cpath dirn = discrimination_pf env sigma e (t,t1,t2) discriminator lbeq in let pf_ty = mkArrow eqn absurd_term in let absurd_clause = apply_on_clause (pf,pf_ty) eq_clause in + let absurd_term = EConstr.of_constr absurd_term in let pf = Clenvtac.clenv_value_cast_meta absurd_clause in Proofview.Unsafe.tclEVARS sigma <*> Proofview.tclEFFECTS eff <*> @@ -1023,18 +1029,15 @@ let discrEq (lbeq,_,(t,t1,t2) as u) eq_clause = let onEquality with_evars tac (c,lbindc) = Proofview.Goal.nf_enter { enter = begin fun gl -> - let c = EConstr.of_constr c in - let lbindc = Miscops.map_bindings EConstr.of_constr lbindc in let type_of = pf_unsafe_type_of gl in let reduce_to_quantified_ind = pf_apply Tacred.reduce_to_quantified_ind gl in let t = type_of c in - let t' = try snd (reduce_to_quantified_ind (EConstr.of_constr t)) with UserError _ -> t in - let t' = EConstr.of_constr t' in + let t = EConstr.of_constr t in + let t' = try snd (reduce_to_quantified_ind t) with UserError _ -> t in let eq_clause = pf_apply make_clenv_binding gl (c,t') lbindc in let eq_clause' = Clenvtac.clenv_pose_dependent_evars with_evars eq_clause in let eqn = clenv_type eq_clause' in - let eqn = EConstr.Unsafe.to_constr eqn in - let (eq,u,eq_args) = find_this_eq_data_decompose gl (EConstr.of_constr eqn) in + let (eq,u,eq_args) = find_this_eq_data_decompose gl eqn in tclTHEN (Proofview.Unsafe.tclEVARS eq_clause'.evd) (tac (eq,eqn,eq_args) eq_clause') @@ -1049,14 +1052,14 @@ let onNegatedEquality with_evars tac = | Prod (_,t,u) when is_empty_type sigma (EConstr.of_constr u) -> tclTHEN introf (onLastHypId (fun id -> - onEquality with_evars tac (mkVar id,NoBindings))) + onEquality with_evars tac (EConstr.mkVar id,NoBindings))) | _ -> tclZEROMSG (str "Not a negated primitive equality.") end } let discrSimpleClause with_evars = function | None -> onNegatedEquality with_evars discrEq - | Some id -> onEquality with_evars discrEq (mkVar id,NoBindings) + | Some id -> onEquality with_evars discrEq (EConstr.mkVar id,NoBindings) let discr with_evars = onEquality with_evars discrEq @@ -1070,7 +1073,7 @@ let discrEverywhere with_evars = (tclTHEN (tclREPEAT introf) (tryAllHyps - (fun id -> tclCOMPLETE (discr with_evars (mkVar id,NoBindings))))) + (fun id -> tclCOMPLETE (discr with_evars (EConstr.mkVar id,NoBindings))))) else (* <= 8.2 compat *) tryAllHypsAndConcl (discrSimpleClause with_evars)) (* (fun gls -> @@ -1194,17 +1197,15 @@ let sig_clausal_form env sigma sort_of_ty siglen ty dflt = | (_sigS,[a;p]) -> (EConstr.Unsafe.to_constr a, EConstr.Unsafe.to_constr p) | _ -> anomaly ~label:"sig_clausal_form" (Pp.str "should be a sigma type") in let ev = Evarutil.e_new_evar env evdref (EConstr.of_constr a) in - let rty = beta_applist sigma (EConstr.of_constr p_i_minus_1,[EConstr.of_constr ev]) in + let rty = beta_applist sigma (EConstr.of_constr p_i_minus_1,[ev]) in let tuple_tail = sigrec_clausal_form (siglen-1) rty in - match - Evd.existential_opt_value !evdref - (destEvar ev) - with + let evopt = match EConstr.kind !evdref ev with Evar _ -> None | _ -> Some ev in + match evopt with | Some w -> - let w_type = unsafe_type_of env sigma (EConstr.of_constr w) in + let w_type = unsafe_type_of env !evdref w in if Evarconv.e_cumul env evdref (EConstr.of_constr w_type) (EConstr.of_constr a) then let exist_term = Evarutil.evd_comb1 (Evd.fresh_global env) evdref sigdata.intro in - applist(exist_term,[a;p_i_minus_1;w;tuple_tail]) + applist(exist_term,[a;p_i_minus_1;EConstr.Unsafe.to_constr w;tuple_tail]) else error "Cannot solve a unification problem." | None -> @@ -1354,7 +1355,8 @@ let inject_if_homogenous_dependent_pair ty = [Proofview.tclEFFECTS eff; intro; onLastHyp (fun hyp -> - tclTHENS (cut (mkApp (ceq,new_eq_args))) + let hyp = EConstr.Unsafe.to_constr hyp in + tclTHENS (cut (EConstr.of_constr (mkApp (ceq,new_eq_args)))) [clear [destVar hyp]; Proofview.V82.tactic (Tacmach.refine (EConstr.of_constr (mkApp(inj2,[|ar1.(0);mkConst c;ar1.(1);ar1.(2);ar1.(3);ar2.(3);hyp|])))) @@ -1404,7 +1406,7 @@ let inject_at_positions env sigma l2r (eq,_,(t,t1,t2)) eq_clause posns tac = Proofview.tclTHEN (Proofview.Unsafe.tclEVARS !evdref) (Tacticals.New.tclTHENFIRST (Proofview.tclIGNORE (Proofview.Monad.List.map - (fun (pf,ty) -> tclTHENS (cut ty) + (fun (pf,ty) -> tclTHENS (cut (EConstr.of_constr ty)) [inject_if_homogenous_dependent_pair (EConstr.of_constr ty); Proofview.V82.tactic (Tacmach.refine (EConstr.of_constr pf))]) (if l2r then List.rev injectors else injectors))) @@ -1452,6 +1454,7 @@ let injEq ?(old=false) with_evars clear_flag ipats = let destopt = match kind_of_term c with | Var id -> get_previous_hyp_position id gl | _ -> MoveLast in + let c = EConstr.of_constr c in let clear_tac = tclTRY (apply_clear_request clear_flag dft_clear_flag c) in (* Try should be removal if dependency were treated *) @@ -1497,12 +1500,11 @@ let dEqThen with_evars ntac = function let dEq with_evars = dEqThen with_evars (fun clear_flag c x -> + let c = EConstr.of_constr c in (apply_clear_request clear_flag (use_clear_hyp_by_default ()) c)) let intro_decomp_eq tac data (c, t) = Proofview.Goal.enter { enter = begin fun gl -> - let c = EConstr.of_constr c in - let t = EConstr.of_constr t in let cl = pf_apply make_clenv_binding gl (c, t) NoBindings in decompEqThen (fun _ -> tac) data cl end } @@ -1596,13 +1598,16 @@ let subst_tuple_term env sigma dep_pair1 dep_pair2 b = (* on for further iterated sigma-tuples *) let cutSubstInConcl l2r eqn = + let eqn = EConstr.of_constr eqn in Proofview.Goal.nf_s_enter { s_enter = begin fun gl -> let env = Proofview.Goal.env gl in let sigma = Proofview.Goal.sigma gl in - let (lbeq,u,(t,e1,e2)) = find_eq_data_decompose gl (EConstr.of_constr eqn) in + let (lbeq,u,(t,e1,e2)) = find_eq_data_decompose gl eqn in let typ = pf_concl gl in let (e1,e2) = if l2r then (e1,e2) else (e2,e1) in let Sigma ((typ, expected), sigma, p) = subst_tuple_term env sigma e1 e2 typ in + let typ = EConstr.of_constr typ in + let expected = EConstr.of_constr expected in let tac = tclTHENFIRST (tclTHENLIST [ @@ -1615,13 +1620,16 @@ let cutSubstInConcl l2r eqn = end } let cutSubstInHyp l2r eqn id = + let eqn = EConstr.of_constr eqn in Proofview.Goal.nf_s_enter { s_enter = begin fun gl -> let env = Proofview.Goal.env gl in let sigma = Proofview.Goal.sigma gl in - let (lbeq,u,(t,e1,e2)) = find_eq_data_decompose gl (EConstr.of_constr eqn) in + let (lbeq,u,(t,e1,e2)) = find_eq_data_decompose gl eqn in let typ = pf_get_hyp_typ id gl in let (e1,e2) = if l2r then (e1,e2) else (e2,e1) in let Sigma ((typ, expected), sigma, p) = subst_tuple_term env sigma e1 e2 typ in + let typ = EConstr.of_constr typ in + let expected = EConstr.of_constr expected in let tac = tclTHENFIRST (tclTHENLIST [ @@ -1653,8 +1661,9 @@ let cutRewriteInHyp l2r eqn id = cutRewriteClause l2r eqn (Some id) let cutRewriteInConcl l2r eqn = cutRewriteClause l2r eqn None let substClause l2r c cls = + let c = EConstr.of_constr c in Proofview.Goal.enter { enter = begin fun gl -> - let eq = pf_apply get_type_of gl (EConstr.of_constr c) in + let eq = pf_apply get_type_of gl c in tclTHENS (cutSubstClause l2r eq cls) [Proofview.tclUNIT (); exact_no_check c] end } @@ -1937,7 +1946,7 @@ let replace_term dir_opt c = (* Declare rewriting tactic for intro patterns "<-" and "->" *) let _ = - let gmr l2r with_evars tac c = general_rewrite_clause l2r with_evars tac c in + let gmr l2r with_evars tac c = general_rewrite_clause l2r with_evars (Miscops.map_with_bindings EConstr.Unsafe.to_constr tac) c in Hook.set Tactics.general_rewrite_clause gmr let _ = Hook.set Tactics.subst_one subst_one diff --git a/tactics/equality.mli b/tactics/equality.mli index 779d1e9b2..97f51ae20 100644 --- a/tactics/equality.mli +++ b/tactics/equality.mli @@ -60,30 +60,30 @@ val general_rewrite_clause : orientation -> evars_flag -> ?tac:(unit Proofview.tactic * conditions) -> constr with_bindings -> clause -> unit Proofview.tactic val general_multi_rewrite : - evars_flag -> (bool * multi * clear_flag * delayed_open_constr_with_bindings) list -> + evars_flag -> (bool * multi * clear_flag * EConstr.constr with_bindings delayed_open) list -> clause -> (unit Proofview.tactic * conditions) option -> unit Proofview.tactic val replace_in_clause_maybe_by : constr -> constr -> clause -> unit Proofview.tactic option -> unit Proofview.tactic val replace : constr -> constr -> unit Proofview.tactic val replace_by : constr -> constr -> unit Proofview.tactic -> unit Proofview.tactic -val discr : evars_flag -> constr with_bindings -> unit Proofview.tactic +val discr : evars_flag -> EConstr.constr with_bindings -> unit Proofview.tactic val discrConcl : unit Proofview.tactic val discrHyp : Id.t -> unit Proofview.tactic val discrEverywhere : evars_flag -> unit Proofview.tactic val discr_tac : evars_flag -> - constr with_bindings destruction_arg option -> unit Proofview.tactic + EConstr.constr with_bindings destruction_arg option -> unit Proofview.tactic val inj : intro_patterns option -> evars_flag -> - clear_flag -> constr with_bindings -> unit Proofview.tactic + clear_flag -> EConstr.constr with_bindings -> unit Proofview.tactic val injClause : intro_patterns option -> evars_flag -> - constr with_bindings destruction_arg option -> unit Proofview.tactic + EConstr.constr with_bindings destruction_arg option -> unit Proofview.tactic val injHyp : clear_flag -> Id.t -> unit Proofview.tactic val injConcl : unit Proofview.tactic val simpleInjClause : evars_flag -> - constr with_bindings destruction_arg option -> unit Proofview.tactic + EConstr.constr with_bindings destruction_arg option -> unit Proofview.tactic -val dEq : evars_flag -> constr with_bindings destruction_arg option -> unit Proofview.tactic -val dEqThen : evars_flag -> (clear_flag -> constr -> int -> unit Proofview.tactic) -> constr with_bindings destruction_arg option -> unit Proofview.tactic +val dEq : evars_flag -> EConstr.constr with_bindings destruction_arg option -> unit Proofview.tactic +val dEqThen : evars_flag -> (clear_flag -> constr -> int -> unit Proofview.tactic) -> EConstr.constr with_bindings destruction_arg option -> unit Proofview.tactic val make_iterated_tuple : env -> evar_map -> constr -> (constr * types) -> evar_map * (constr * constr * constr) diff --git a/tactics/hints.ml b/tactics/hints.ml index ea95fb1ad..560e7e43d 100644 --- a/tactics/hints.ml +++ b/tactics/hints.ml @@ -1320,7 +1320,7 @@ let make_local_hint_db env sigma ts eapply lems = let map c = let sigma = Sigma.Unsafe.of_evar_map sigma in let Sigma (c, sigma, _) = c.delayed env sigma in - (Sigma.to_evar_map sigma, c) + (Sigma.to_evar_map sigma, EConstr.Unsafe.to_constr c) in let lems = List.map map lems in let sign = Environ.named_context env in diff --git a/tactics/inv.ml b/tactics/inv.ml index a971b9356..c66b356c7 100644 --- a/tactics/inv.ml +++ b/tactics/inv.ml @@ -284,7 +284,7 @@ let error_too_many_names pats = tclZEROMSG ~loc ( str "Unexpected " ++ str (String.plural (List.length pats) "introduction pattern") ++ - str ": " ++ pr_enum (Miscprint.pr_intro_pattern (fun c -> Printer.pr_constr (fst (run_delayed env Evd.empty c)))) pats ++ + str ": " ++ pr_enum (Miscprint.pr_intro_pattern (fun c -> Printer.pr_constr (EConstr.Unsafe.to_constr (fst (run_delayed env Evd.empty c))))) pats ++ str ".") let get_names (allow_conj,issimple) (loc, pat as x) = match pat with @@ -369,7 +369,7 @@ let projectAndApply as_mode thin avoid id eqname names depids = (* and apply a trailer which again try to substitute *) (fun id -> dEqThen false (deq_trailer id) - (Some (None,ElimOnConstr (mkVar id,NoBindings)))) + (Some (None,ElimOnConstr (EConstr.mkVar id,NoBindings)))) id let nLastDecls i tac = @@ -443,7 +443,6 @@ let raw_inversion inv_kind id status names = let msg = str "The type of " ++ pr_id id ++ str " is not inductive." in CErrors.user_err msg in - let t = EConstr.of_constr t in let IndType (indf,realargs) = find_rectype env sigma t in let evdref = ref sigma in let (elim_predicate, args) = @@ -457,6 +456,7 @@ let raw_inversion inv_kind id status names = Reduction.beta_appvect elim_predicate (Array.of_list realargs), case_nodep_then_using in + let cut_concl = EConstr.of_constr cut_concl in let refined id = let prf = mkApp (mkVar id, args) in let prf = EConstr.of_constr prf in @@ -505,7 +505,7 @@ let inv k = inv_gen k NoDep let inv_tac id = inv FullInversion None (NamedHyp id) let inv_clear_tac id = inv FullInversionClear None (NamedHyp id) -let dinv k c = inv_gen k (Dep c) +let dinv k c = inv_gen k (Dep (Option.map EConstr.Unsafe.to_constr c)) let dinv_tac id = dinv FullInversion None None (NamedHyp id) let dinv_clear_tac id = dinv FullInversionClear None None (NamedHyp id) diff --git a/tactics/inv.mli b/tactics/inv.mli index df629e7c9..6bb2b7282 100644 --- a/tactics/inv.mli +++ b/tactics/inv.mli @@ -20,7 +20,7 @@ val inv_clause : val inv : inversion_kind -> or_and_intro_pattern option -> quantified_hypothesis -> unit Proofview.tactic -val dinv : inversion_kind -> constr option -> +val dinv : inversion_kind -> EConstr.constr option -> or_and_intro_pattern option -> quantified_hypothesis -> unit Proofview.tactic val inv_tac : Id.t -> unit Proofview.tactic diff --git a/tactics/leminv.ml b/tactics/leminv.ml index 16a048af8..a94238418 100644 --- a/tactics/leminv.ml +++ b/tactics/leminv.ml @@ -291,5 +291,5 @@ let lemInvIn id c ids = let lemInvIn_gen id c l = try_intros_until (fun id -> lemInvIn id c l) id let lemInv_clause id c = function - | [] -> lemInv_gen id c - | l -> lemInvIn_gen id c l + | [] -> lemInv_gen id (EConstr.Unsafe.to_constr c) + | l -> lemInvIn_gen id (EConstr.Unsafe.to_constr c) l diff --git a/tactics/leminv.mli b/tactics/leminv.mli index c6ed9606f..58b82002d 100644 --- a/tactics/leminv.mli +++ b/tactics/leminv.mli @@ -12,7 +12,7 @@ open Constrexpr open Misctypes val lemInv_clause : - quantified_hypothesis -> constr -> Id.t list -> unit Proofview.tactic + quantified_hypothesis -> EConstr.constr -> Id.t list -> unit Proofview.tactic val add_inversion_lemma_exn : Id.t -> constr_expr -> glob_sort -> bool -> (Id.t -> unit Proofview.tactic) -> diff --git a/tactics/tacticals.ml b/tactics/tacticals.ml index 0546132c1..e15ee149d 100644 --- a/tactics/tacticals.ml +++ b/tactics/tacticals.ml @@ -73,7 +73,7 @@ let nthDecl m gl = with Failure _ -> error "No such assumption." let nthHypId m gl = nthDecl m gl |> NamedDecl.get_id -let nthHyp m gl = mkVar (nthHypId m gl) +let nthHyp m gl = EConstr.mkVar (nthHypId m gl) let lastDecl gl = nthDecl 1 gl let lastHypId gl = nthHypId 1 gl @@ -564,7 +564,7 @@ module New = struct let gl = Proofview.Goal.assume gl in nthDecl m gl |> NamedDecl.get_id let nthHyp m gl = - mkVar (nthHypId m gl) + EConstr.mkVar (nthHypId m gl) let onNthHypId m tac = Proofview.Goal.enter { enter = begin fun gl -> tac (nthHypId m gl) end } @@ -680,7 +680,6 @@ module New = struct let elimination_then tac c = Proofview.Goal.nf_enter { enter = begin fun gl -> let (ind,t) = pf_reduce_to_quantified_ind gl (EConstr.of_constr (pf_unsafe_type_of gl c)) in - let t = EConstr.of_constr t in let isrec,mkelim = match (Global.lookup_mind (fst (fst ind))).mind_record with | None -> true,gl_make_elim diff --git a/tactics/tacticals.mli b/tactics/tacticals.mli index 974bf83a3..2c3e51280 100644 --- a/tactics/tacticals.mli +++ b/tactics/tacticals.mli @@ -58,17 +58,17 @@ val tclIFTHENTRYELSEMUST : tactic -> tactic -> tactic (** {6 Tacticals applying to hypotheses } *) val onNthHypId : int -> (Id.t -> tactic) -> tactic -val onNthHyp : int -> (constr -> tactic) -> tactic +val onNthHyp : int -> (EConstr.constr -> tactic) -> tactic val onNthDecl : int -> (Context.Named.Declaration.t -> tactic) -> tactic val onLastHypId : (Id.t -> tactic) -> tactic -val onLastHyp : (constr -> tactic) -> tactic +val onLastHyp : (EConstr.constr -> tactic) -> tactic val onLastDecl : (Context.Named.Declaration.t -> tactic) -> tactic val onNLastHypsId : int -> (Id.t list -> tactic) -> tactic val onNLastHyps : int -> (constr list -> tactic) -> tactic val onNLastDecls : int -> (Context.Named.t -> tactic) -> tactic val lastHypId : goal sigma -> Id.t -val lastHyp : goal sigma -> constr +val lastHyp : goal sigma -> EConstr.constr val lastDecl : goal sigma -> Context.Named.Declaration.t val nLastHypsId : int -> goal sigma -> Id.t list val nLastHyps : int -> goal sigma -> constr list @@ -236,7 +236,7 @@ module New : sig val onNthHypId : int -> (identifier -> unit tactic) -> unit tactic val onLastHypId : (identifier -> unit tactic) -> unit tactic - val onLastHyp : (constr -> unit tactic) -> unit tactic + val onLastHyp : (EConstr.constr -> unit tactic) -> unit tactic val onLastDecl : (Context.Named.Declaration.t -> unit tactic) -> unit tactic val onHyps : ([ `NF ], Context.Named.t) Proofview.Goal.enter -> diff --git a/tactics/tactics.ml b/tactics/tactics.ml index a04fb7ca2..b9da11021 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -6,14 +6,17 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +module CVars = Vars + open Pp open CErrors open Util open Names open Nameops open Term -open Vars open Termops +open EConstr +open Vars open Find_subterm open Namegen open Declarations @@ -48,7 +51,7 @@ open Context.Named.Declaration module RelDecl = Context.Rel.Declaration module NamedDecl = Context.Named.Declaration -let inj_with_occurrences e = (AllOccurrences,e) +let inj_with_occurrences e = (AllOccurrences,EConstr.Unsafe.to_constr e) let dloc = Loc.ghost @@ -167,6 +170,26 @@ let _ = (* Primitive tactics *) (******************************************) +let local_assum (na, t) = + let open Context.Rel.Declaration in + let inj = EConstr.Unsafe.to_constr in + LocalAssum (na, inj t) + +let local_def (na, b, t) = + let open Context.Rel.Declaration in + let inj = EConstr.Unsafe.to_constr in + LocalDef (na, inj b, inj t) + +let nlocal_assum (na, t) = + let open Context.Named.Declaration in + let inj = EConstr.Unsafe.to_constr in + LocalAssum (na, inj t) + +let nlocal_def (na, b, t) = + let open Context.Named.Declaration in + let inj = EConstr.Unsafe.to_constr in + LocalDef (na, inj b, inj t) + (** This tactic creates a partial proof realizing the introduction rule, but does not check anything. *) let unsafe_intro env store decl b = @@ -176,14 +199,15 @@ let unsafe_intro env store decl b = let inst = List.map (NamedDecl.get_id %> mkVar) (named_context env) in let ninst = mkRel 1 :: inst in let nb = subst1 (mkVar (NamedDecl.get_id decl)) b in - let Sigma (ev, sigma, p) = new_evar_instance nctx sigma (EConstr.of_constr nb) ~principal:true ~store ninst in - Sigma (EConstr.of_constr (mkNamedLambda_or_LetIn decl ev), sigma, p) + let Sigma (ev, sigma, p) = new_evar_instance nctx sigma nb ~principal:true ~store ninst in + Sigma (mkNamedLambda_or_LetIn decl ev, sigma, p) end } let introduction ?(check=true) id = Proofview.Goal.enter { enter = begin fun gl -> let gl = Proofview.Goal.assume gl in let concl = Proofview.Goal.concl gl in + let concl = EConstr.of_constr concl in let sigma = Tacmach.New.project gl in let hyps = named_context_val (Proofview.Goal.env gl) in let store = Proofview.Goal.extra gl in @@ -193,9 +217,9 @@ let introduction ?(check=true) id = (str "Variable " ++ pr_id id ++ str " is already declared.") in let open Context.Named.Declaration in - match kind_of_term (whd_evar sigma concl) with - | Prod (_, t, b) -> unsafe_intro env store (LocalAssum (id, t)) b - | LetIn (_, c, t, b) -> unsafe_intro env store (LocalDef (id, c, t)) b + match EConstr.kind sigma concl with + | Prod (_, t, b) -> unsafe_intro env store (nlocal_assum (id, t)) b + | LetIn (_, c, t, b) -> unsafe_intro env store (nlocal_def (id, c, t)) b | _ -> raise (RefinerError IntroNeedsProduct) end } @@ -206,19 +230,19 @@ let convert_concl ?(check=true) ty k = let env = Proofview.Goal.env gl in let store = Proofview.Goal.extra gl in let conclty = Proofview.Goal.raw_concl gl in - let ty = EConstr.of_constr ty in + let conclty = EConstr.of_constr conclty in Refine.refine ~unsafe:true { run = begin fun sigma -> let Sigma ((), sigma, p) = if check then begin let sigma = Sigma.to_evar_map sigma in ignore (Typing.unsafe_type_of env sigma ty); - let sigma,b = Reductionops.infer_conv env sigma ty (EConstr.of_constr conclty) in + let sigma,b = Reductionops.infer_conv env sigma ty conclty in if not b then error "Not convertible."; Sigma.Unsafe.of_pair ((), sigma) end else Sigma.here () sigma in let Sigma (x, sigma, q) = Evarutil.new_evar env sigma ~principal:true ~store ty in let ans = if k == DEFAULTcast then x else mkCast(x,k,conclty) in - Sigma (EConstr.of_constr ans, sigma, p +> q) + Sigma (ans, sigma, p +> q) end } end } @@ -227,12 +251,12 @@ let convert_hyp ?(check=true) d = let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in let ty = Proofview.Goal.raw_concl gl in + let ty = EConstr.of_constr ty in let store = Proofview.Goal.extra gl in let sign = convert_hyp check (named_context_val env) sigma d in let env = reset_with_named_context sign env in Refine.refine ~unsafe:true { run = begin fun sigma -> - let Sigma (c, sigma, p) = Evarutil.new_evar env sigma ~principal:true ~store (EConstr.of_constr ty) in - Sigma (EConstr.of_constr c, sigma, p) + Evarutil.new_evar env sigma ~principal:true ~store ty end } end } @@ -250,8 +274,8 @@ let convert_gen pb x y = Tacticals.New.tclFAIL 0 (str "Not convertible") end } -let convert x y = convert_gen Reduction.CONV (EConstr.of_constr x) (EConstr.of_constr y) -let convert_leq x y = convert_gen Reduction.CUMUL (EConstr.of_constr x) (EConstr.of_constr y) +let convert x y = convert_gen Reduction.CONV x y +let convert_leq x y = convert_gen Reduction.CUMUL x y let clear_dependency_msg env sigma id = function | Evarutil.OccurHypInSimpleClause None -> @@ -300,10 +324,10 @@ let clear_gen fail = function try clear_hyps_in_evi env evdref (named_context_val env) concl ids with Evarutil.ClearDependencyError (id,err) -> fail env sigma id err in + let concl = EConstr.of_constr concl in let env = reset_with_named_context hyps env in let tac = Refine.refine ~unsafe:true { run = fun sigma -> - let Sigma (c, sigma, p) = Evarutil.new_evar env sigma ~principal:true (EConstr.of_constr concl) in - Sigma (EConstr.of_constr c, sigma, p) + Evarutil.new_evar env sigma ~principal:true concl } in Sigma.Unsafe.of_pair (tac, !evdref) end } @@ -312,14 +336,15 @@ let clear ids = clear_gen error_clear_dependency ids let clear_for_replacing ids = clear_gen error_replacing_dependency ids let apply_clear_request clear_flag dft c = + Proofview.tclEVARMAP >>= fun sigma -> let check_isvar c = - if not (isVar c) then + if not (isVar sigma c) then error "keep/clear modifiers apply only to hypothesis names." in let doclear = match clear_flag with - | None -> dft && isVar c + | None -> dft && isVar sigma c | Some true -> check_isvar c; true | Some false -> false in - if doclear then clear [destVar c] + if doclear then clear [destVar sigma c] else Tacticals.New.tclIDTAC (* Moving hypotheses *) @@ -328,13 +353,13 @@ let move_hyp id dest = let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in let ty = Proofview.Goal.raw_concl gl in + let ty = EConstr.of_constr ty in let store = Proofview.Goal.extra gl in let sign = named_context_val env in let sign' = move_hyp_in_named_context sigma id dest sign in let env = reset_with_named_context sign' env in Refine.refine ~unsafe:true { run = begin fun sigma -> - let Sigma (c, sigma, p) = Evarutil.new_evar env sigma ~principal:true ~store (EConstr.of_constr ty) in - Sigma (EConstr.of_constr c, sigma, p) + Evarutil.new_evar env sigma ~principal:true ~store ty end } end } @@ -376,20 +401,20 @@ let rename_hyp repl = with Not_found -> () in (** All is well *) - let make_subst (src, dst) = (src, mkVar dst) in + let make_subst (src, dst) = (src, Constr.mkVar dst) in let subst = List.map make_subst repl in - let subst c = Vars.replace_vars subst c in + let subst c = CVars.replace_vars subst c in let map decl = decl |> NamedDecl.map_id (fun id -> try List.assoc_f Id.equal id repl with Not_found -> id) |> NamedDecl.map_constr subst in let nhyps = List.map map hyps in let nconcl = subst concl in + let nconcl = EConstr.of_constr nconcl in let nctx = Environ.val_of_named_context nhyps in let instance = List.map (NamedDecl.get_id %> mkVar) hyps in Refine.refine ~unsafe:true { run = begin fun sigma -> - let Sigma (c, sigma, p) = Evarutil.new_evar_instance nctx sigma (EConstr.of_constr nconcl) ~principal:true ~store instance in - Sigma (EConstr.of_constr c, sigma, p) + Evarutil.new_evar_instance nctx sigma nconcl ~principal:true ~store instance end } end } @@ -456,8 +481,7 @@ let find_name mayrepl decl naming gl = match naming with let assert_before_then_gen b naming t tac = let open Context.Rel.Declaration in Proofview.Goal.enter { enter = begin fun gl -> - let id = find_name b (LocalAssum (Anonymous,t)) naming gl in - let t = EConstr.of_constr t in + let id = find_name b (local_assum (Anonymous,t)) naming gl in Tacticals.New.tclTHENLAST (Proofview.V82.tactic (fun gl -> @@ -476,8 +500,7 @@ let assert_before_replacing id = assert_before_gen true (NamingMustBe (dloc,id)) let assert_after_then_gen b naming t tac = let open Context.Rel.Declaration in Proofview.Goal.enter { enter = begin fun gl -> - let id = find_name b (LocalAssum (Anonymous,t)) naming gl in - let t = EConstr.of_constr t in + let id = find_name b (local_assum (Anonymous,t)) naming gl in Tacticals.New.tclTHENFIRST (Proofview.V82.tactic (fun gl -> @@ -501,20 +524,20 @@ let rec mk_holes : type r s. _ -> r Sigma.t -> (s, r) Sigma.le -> _ -> (_, s) Si fun env sigma p -> function | [] -> Sigma ([], sigma, p) | arg :: rem -> - let Sigma (arg, sigma, q) = Evarutil.new_evar env sigma (EConstr.of_constr arg) in + let Sigma (arg, sigma, q) = Evarutil.new_evar env sigma arg in let Sigma (rem, sigma, r) = mk_holes env sigma (p +> q) rem in Sigma (arg :: rem, sigma, r) -let rec check_mutind env sigma k cl = match kind_of_term (strip_outer_cast sigma (EConstr.of_constr cl)) with +let rec check_mutind env sigma k cl = match EConstr.kind sigma (EConstr.of_constr (strip_outer_cast sigma cl)) with | Prod (na, c1, b) -> if Int.equal k 1 then try - let ((sp, _), u), _ = find_inductive env sigma (EConstr.of_constr c1) in + let ((sp, _), u), _ = find_inductive env sigma c1 in (sp, u) with Not_found -> error "Cannot do a fixpoint on a non inductive type." else let open Context.Rel.Declaration in - check_mutind (push_rel (LocalAssum (na, c1)) env) sigma (pred k) b + check_mutind (push_rel (local_assum (na, c1)) env) sigma (pred k) b | _ -> error "Not enough products." (* Refine as a fixpoint *) @@ -522,20 +545,20 @@ let mutual_fix f n rest j = Proofview.Goal.nf_enter { enter = begin fun gl -> let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in let concl = Proofview.Goal.concl gl in + let concl = EConstr.of_constr concl in let (sp, u) = check_mutind env sigma n concl in let firsts, lasts = List.chop j rest in let all = firsts @ (f, n, concl) :: lasts in let rec mk_sign sign = function | [] -> sign | (f, n, ar) :: oth -> - let open Context.Named.Declaration in let (sp', u') = check_mutind env sigma n ar in if not (eq_mind sp sp') then error "Fixpoints should be on the same mutual inductive declaration."; if mem_named_context_val f sign then user_err ~hdr:"Logic.prim_refiner" (str "Name " ++ pr_id f ++ str " already used in the environment"); - mk_sign (push_named_context_val (LocalAssum (f, ar)) sign) oth + mk_sign (push_named_context_val (nlocal_assum (f, ar)) sign) oth in let nenv = reset_with_named_context (mk_sign (named_context_val env) all) env in Refine.refine { run = begin fun sigma -> @@ -546,8 +569,7 @@ let mutual_fix f n rest j = Proofview.Goal.nf_enter { enter = begin fun gl -> let funnames = Array.of_list (List.map (fun i -> Name i) ids) in let typarray = Array.of_list (List.map pi3 all) in let bodies = Array.of_list evs in - let oterm = Term.mkFix ((indxs,0),(funnames,typarray,bodies)) in - let oterm = EConstr.of_constr oterm in + let oterm = mkFix ((indxs,0),(funnames,typarray,bodies)) in Sigma (oterm, sigma, p) end } end } @@ -563,14 +585,14 @@ let fix ido n = match ido with mutual_fix id n [] 0 let rec check_is_mutcoind env sigma cl = - let b = whd_all env sigma (EConstr.of_constr cl) in - match kind_of_term b with + let b = whd_all env sigma cl in + let b = EConstr.of_constr b in + match EConstr.kind sigma b with | Prod (na, c1, b) -> - let open Context.Rel.Declaration in - check_is_mutcoind (push_rel (LocalAssum (na,c1)) env) sigma b + check_is_mutcoind (push_rel (local_assum (na,c1)) env) sigma b | _ -> try - let _ = find_coinductive env sigma (EConstr.of_constr b) in () + let _ = find_coinductive env sigma b in () with Not_found -> error "All methods must construct elements in coinductive types." @@ -579,16 +601,16 @@ let mutual_cofix f others j = Proofview.Goal.nf_enter { enter = begin fun gl -> let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in let concl = Proofview.Goal.concl gl in + let concl = EConstr.of_constr concl in let firsts,lasts = List.chop j others in let all = firsts @ (f, concl) :: lasts in List.iter (fun (_, c) -> check_is_mutcoind env sigma c) all; let rec mk_sign sign = function | [] -> sign | (f, ar) :: oth -> - let open Context.Named.Declaration in if mem_named_context_val f sign then error "Name already used in the environment."; - mk_sign (push_named_context_val (LocalAssum (f, ar)) sign) oth + mk_sign (push_named_context_val (nlocal_assum (f, ar)) sign) oth in let nenv = reset_with_named_context (mk_sign (named_context_val env) all) env in Refine.refine { run = begin fun sigma -> @@ -598,8 +620,7 @@ let mutual_cofix f others j = Proofview.Goal.nf_enter { enter = begin fun gl -> let funnames = Array.of_list (List.map (fun i -> Name i) ids) in let typarray = Array.of_list types in let bodies = Array.of_list evs in - let oterm = Term.mkCoFix (0, (funnames, typarray, bodies)) in - let oterm = EConstr.of_constr oterm in + let oterm = mkCoFix (0, (funnames, typarray, bodies)) in Sigma (oterm, sigma, p) end } end } @@ -618,20 +639,23 @@ let cofix ido = match ido with (* Reduction and conversion tactics *) (**************************************************************) -type tactic_reduction = env -> evar_map -> EConstr.t -> constr +type tactic_reduction = env -> evar_map -> constr -> Constr.constr let pf_reduce_decl redfun where decl gl = let open Context.Named.Declaration in - let redfun' c = Tacmach.New.pf_apply redfun gl (EConstr.of_constr c) in + let redfun' c = EConstr.of_constr (Tacmach.New.pf_apply redfun gl c) in match decl with | LocalAssum (id,ty) -> + let ty = EConstr.of_constr ty in if where == InHypValueOnly then user_err (pr_id id ++ str " has no value."); - LocalAssum (id,redfun' ty) + nlocal_assum (id,redfun' ty) | LocalDef (id,b,ty) -> + let b = EConstr.of_constr b in + let ty = EConstr.of_constr ty in let b' = if where != InHypTypeOnly then redfun' b else b in let ty' = if where != InHypValueOnly then redfun' ty else ty in - LocalDef (id,b',ty') + nlocal_def (id,b',ty') (* Possibly equip a reduction with the occurrences mentioned in an occurrence clause *) @@ -703,7 +727,7 @@ let bind_red_expr_occurrences occs nbcl redexp = let reduct_in_concl (redfun,sty) = Proofview.Goal.nf_enter { enter = begin fun gl -> - convert_concl_no_check (Tacmach.New.pf_apply redfun gl (EConstr.of_constr (Tacmach.New.pf_concl gl))) sty + convert_concl_no_check (EConstr.of_constr (Tacmach.New.pf_apply redfun gl (EConstr.of_constr (Tacmach.New.pf_concl gl)))) sty end } let reduct_in_hyp ?(check=false) redfun (id,where) = @@ -739,6 +763,7 @@ let e_reduct_in_concl ~check (redfun, sty) = Proofview.Goal.nf_s_enter { s_enter = begin fun gl -> let sigma = Proofview.Goal.sigma gl in let Sigma (c', sigma, p) = redfun.e_redfun (Tacmach.New.pf_env gl) sigma (EConstr.of_constr (Tacmach.New.pf_concl gl)) in + let c' = EConstr.of_constr c' in Sigma (convert_concl ~check c' sty, sigma, p) end } @@ -759,6 +784,7 @@ let e_change_in_concl (redfun,sty) = Proofview.Goal.s_enter { s_enter = begin fun gl -> let sigma = Proofview.Goal.sigma gl in let Sigma (c, sigma, p) = redfun.e_redfun (Proofview.Goal.env gl) sigma (EConstr.of_constr (Proofview.Goal.raw_concl gl)) in + let c = EConstr.of_constr c in Sigma (convert_concl_no_check c sty, sigma, p) end } @@ -787,9 +813,10 @@ let e_change_in_hyp redfun (id,where) = Sigma (convert_hyp c, sigma, p) end } -type change_arg = Pattern.patvar_map -> constr Sigma.run +type change_arg = Pattern.patvar_map -> EConstr.constr Sigma.run let make_change_arg c pats = + let pats = Id.Map.map EConstr.of_constr pats in { run = fun sigma -> Sigma.here (replace_vars (Id.Map.bindings pats) c) sigma } let check_types env sigma mayneedglobalcheck deep newc origc = @@ -803,15 +830,15 @@ let check_types env sigma mayneedglobalcheck deep newc origc = let sigma, b = infer_conv ~pb:Reduction.CUMUL env sigma t1 t2 in if not b then if - isSort (whd_all env sigma t1) && - isSort (whd_all env sigma t2) + isSort sigma (EConstr.of_constr (whd_all env sigma t1)) && + isSort sigma (EConstr.of_constr (whd_all env sigma t2)) then (mayneedglobalcheck := true; sigma) else user_err ~hdr:"convert-check-hyp" (str "Types are incompatible.") else sigma end else - if not (isSort (whd_all env sigma t1)) then + if not (isSort sigma (EConstr.of_constr (whd_all env sigma t1))) then user_err ~hdr:"convert-check-hyp" (str "Not a type.") else sigma @@ -819,7 +846,6 @@ let check_types env sigma mayneedglobalcheck deep newc origc = let change_and_check cv_pb mayneedglobalcheck deep t = { e_redfun = begin fun env sigma c -> let Sigma (t', sigma, p) = t.run sigma in let sigma = Sigma.to_evar_map sigma in - let t' = EConstr.of_constr t' in let sigma = check_types env sigma mayneedglobalcheck deep t' c in let sigma, b = infer_conv ~pb:cv_pb env sigma t' c in if not b then user_err ~hdr:"convert-check-hyp" (str "Not convertible."); @@ -886,7 +912,7 @@ let normalise_vm_in_concl = reduct_in_concl (Redexpr.cbv_vm,VMcast) let unfold_in_concl loccname = reduct_in_concl (unfoldn loccname,REVERTcast) let unfold_in_hyp loccname = reduct_in_hyp (unfoldn loccname) let unfold_option loccname = reduct_option (unfoldn loccname,DEFAULTcast) -let pattern_option l = e_reduct_option (pattern_occs (List.map (on_snd EConstr.of_constr) l),DEFAULTcast) +let pattern_option l = e_reduct_option (pattern_occs l,DEFAULTcast) (* The main reduction function *) @@ -951,13 +977,13 @@ let rec intro_then_gen name_flag move_flag force_flag dep_flag tac = Proofview.Goal.enter { enter = begin fun gl -> let sigma = Tacmach.New.project gl in let concl = Proofview.Goal.concl (Proofview.Goal.assume gl) in - let concl = nf_evar (Tacmach.New.project gl) concl in - match kind_of_term concl with - | Prod (name,t,u) when not dep_flag || not (EConstr.Vars.noccurn sigma 1 (EConstr.of_constr u)) -> - let name = find_name false (LocalAssum (name,t)) name_flag gl in + let concl = EConstr.of_constr concl in + match EConstr.kind sigma concl with + | Prod (name,t,u) when not dep_flag || not (noccurn sigma 1 u) -> + let name = find_name false (local_assum (name,t)) name_flag gl in build_intro_tac name move_flag tac - | LetIn (name,b,t,u) when not dep_flag || not (EConstr.Vars.noccurn sigma 1 (EConstr.of_constr u)) -> - let name = find_name false (LocalDef (name,b,t)) name_flag gl in + | LetIn (name,b,t,u) when not dep_flag || not (noccurn sigma 1 u) -> + let name = find_name false (local_def (name,b,t)) name_flag gl in build_intro_tac name move_flag tac | _ -> begin if not force_flag then Proofview.tclZERO (RefinerError IntroNeedsProduct) @@ -1212,12 +1238,10 @@ let map_destruction_arg f sigma = function let finish_delayed_evar_resolution with_evars env sigma f = let ((c, lbind), sigma') = run_delayed env sigma f in - let c = EConstr.of_constr c in let pending = (sigma,sigma') in let sigma' = Sigma.Unsafe.of_evar_map sigma' in let flags = tactic_infer_flags with_evars in let Sigma (c, sigma', _) = finish_evar_resolution ~flags env sigma' (pending,c) in - let c = EConstr.Unsafe.to_constr c in (Sigma.to_evar_map sigma', (c, lbind)) let with_no_bindings (c, lbind) = @@ -1238,12 +1262,15 @@ let cut c = let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in let concl = Tacmach.New.pf_nf_concl gl in + let concl = EConstr.of_constr concl in let is_sort = try (** Backward compat: ensure that [c] is well-typed. *) - let typ = Typing.unsafe_type_of env sigma (EConstr.of_constr c) in - let typ = whd_all env sigma (EConstr.of_constr typ) in - match kind_of_term typ with + let typ = Typing.unsafe_type_of env sigma c in + let typ = EConstr.of_constr typ in + let typ = whd_all env sigma typ in + let typ = EConstr.of_constr typ in + match EConstr.kind sigma typ with | Sort _ -> true | _ -> false with e when Pretype_errors.precatchable_exception e -> false @@ -1251,12 +1278,11 @@ let cut c = if is_sort then let id = next_name_away_with_default "H" Anonymous (Tacmach.New.pf_ids_of_hyps gl) in (** Backward compat: normalize [c]. *) - let c = if normalize_cut then local_strong whd_betaiota sigma (EConstr.of_constr c) else c in + let c = if normalize_cut then EConstr.of_constr (local_strong whd_betaiota sigma c) else c in Refine.refine ~unsafe:true { run = begin fun h -> - let Sigma (f, h, p) = Evarutil.new_evar ~principal:true env h (EConstr.of_constr (mkArrow c (Vars.lift 1 concl))) in - let Sigma (x, h, q) = Evarutil.new_evar env h (EConstr.of_constr c) in + let Sigma (f, h, p) = Evarutil.new_evar ~principal:true env h (mkArrow c (Vars.lift 1 concl)) in + let Sigma (x, h, q) = Evarutil.new_evar env h c in let f = mkLetIn (Name id, x, c, mkApp (Vars.lift 1 f, [|mkRel 1|])) in - let f = EConstr.of_constr f in Sigma (f, h, p +> q) end } else @@ -1264,6 +1290,7 @@ let cut c = end } let error_uninstantiated_metas t clenv = + let t = EConstr.Unsafe.to_constr t in let na = meta_name clenv.evd (List.hd (Metaset.elements (metavars_of t))) in let id = match na with Name id -> id | _ -> anomaly (Pp.str "unnamed dependent meta") in user_err (str "Cannot find an instance for " ++ pr_id id ++ str".") @@ -1276,7 +1303,7 @@ let check_unresolved_evars_of_metas sigma clenv = (match kind_of_term c.rebus with | Evar (evk,_) when Evd.is_undefined clenv.evd evk && not (Evd.mem sigma evk) -> - error_uninstantiated_metas (mkMeta mv) clenv + error_uninstantiated_metas (EConstr.mkMeta mv) clenv | _ -> ()) | _ -> ()) (meta_list clenv.evd) @@ -1301,9 +1328,8 @@ let clenv_refine_in ?(sidecond_first=false) with_evars ?(with_classes=true) else clenv in let new_hyp_typ = clenv_type clenv in - let new_hyp_typ = EConstr.Unsafe.to_constr new_hyp_typ in if not with_evars then check_unresolved_evars_of_metas sigma0 clenv; - if not with_evars && occur_meta clenv.evd (EConstr.of_constr new_hyp_typ) then + if not with_evars && occur_meta clenv.evd new_hyp_typ then error_uninstantiated_metas new_hyp_typ clenv; let new_hyp_prf = clenv_value clenv in let exact_tac = Proofview.V82.tactic (Tacmach.refine_no_check new_hyp_prf) in @@ -1322,22 +1348,22 @@ let clenv_refine_in ?(sidecond_first=false) with_evars ?(with_classes=true) (* Elimination tactics *) (********************************************) -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") -let nth_arg i c = - if Int.equal i (-1) then last_arg c else - match kind_of_term c with +let nth_arg sigma i c = + if Int.equal i (-1) then last_arg sigma c else + match EConstr.kind sigma c with | App (f,cl) -> cl.(i) | _ -> anomaly (Pp.str "nth_arg") -let index_of_ind_arg t = - let rec aux i j t = match kind_of_term t with +let index_of_ind_arg sigma t = + let rec aux i j t = match EConstr.kind sigma t with | Prod (_,t,u) -> (* heuristic *) - if isInd (fst (decompose_app t)) then aux (Some j) (j+1) u + if isInd sigma (fst (decompose_app sigma t)) then aux (Some j) (j+1) u else aux i (j+1) u | _ -> match i with | Some i -> i @@ -1352,30 +1378,31 @@ let enforce_prop_bound_names rename tac = (* so as to avoid having hypothesis such as "t:True", "n:~A" when calling *) (* elim or induction with schemes built by Indrec.build_induction_scheme *) let rec aux env sigma i t = - if i = 0 then t else match kind_of_term t with + if i = 0 then t else match EConstr.kind sigma t with | Prod (Name _ as na,t,t') -> let very_standard = true in let na = - if Retyping.get_sort_family_of env sigma (EConstr.of_constr t) = InProp then + if Retyping.get_sort_family_of env sigma t = InProp then (* "very_standard" says that we should have "H" names only, but this would break compatibility even more... *) - let s = match Namegen.head_name t with + let s = match Namegen.head_name (EConstr.Unsafe.to_constr t) with | Some id when not very_standard -> string_of_id id | _ -> "" in Name (add_suffix Namegen.default_prop_ident s) else na in - mkProd (na,t,aux (push_rel (LocalAssum (na,t)) env) sigma (i-1) t') + mkProd (na,t,aux (push_rel (local_assum (na,t)) env) sigma (i-1) t') | Prod (Anonymous,t,t') -> - mkProd (Anonymous,t,aux (push_rel (LocalAssum (Anonymous,t)) env) sigma (i-1) t') + mkProd (Anonymous,t,aux (push_rel (local_assum (Anonymous,t)) env) sigma (i-1) t') | LetIn (na,c,t,t') -> - mkLetIn (na,c,t,aux (push_rel (LocalDef (na,c,t)) env) sigma (i-1) t') - | _ -> print_int i; Feedback.msg_notice (print_constr t); assert false in + mkLetIn (na,c,t,aux (push_rel (local_def (na,c,t)) env) sigma (i-1) t') + | _ -> assert false in let rename_branch i = Proofview.Goal.nf_enter { enter = begin fun gl -> let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in let t = Proofview.Goal.concl gl in + let t = EConstr.of_constr t in change_concl (aux env sigma i t) end } in (if isrec then Tacticals.New.tclTHENFIRSTn else Tacticals.New.tclTHENLASTn) @@ -1384,10 +1411,10 @@ let enforce_prop_bound_names rename tac = | _ -> tac -let rec contract_letin_in_lam_header c = - match kind_of_term c with - | Lambda (x,t,c) -> mkLambda (x,t,contract_letin_in_lam_header c) - | LetIn (x,b,t,c) -> contract_letin_in_lam_header (subst1 b c) +let rec contract_letin_in_lam_header sigma c = + match EConstr.kind sigma c with + | Lambda (x,t,c) -> mkLambda (x,t,contract_letin_in_lam_header sigma c) + | LetIn (x,b,t,c) -> contract_letin_in_lam_header sigma (subst1 b c) | _ -> c let elimination_clause_scheme with_evars ?(with_classes=true) ?(flags=elim_flags ()) @@ -1395,13 +1422,10 @@ let elimination_clause_scheme with_evars ?(with_classes=true) ?(flags=elim_flags Proofview.Goal.enter { enter = begin fun gl -> let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in - let elim = contract_letin_in_lam_header elim in - let bindings = Miscops.map_bindings EConstr.of_constr bindings in - let elim = EConstr.of_constr elim in - let elimty = EConstr.of_constr elimty in + let elim = contract_letin_in_lam_header sigma elim in let elimclause = make_clenv_binding env sigma (elim, elimty) bindings in let indmv = - (match kind_of_term (nth_arg i (EConstr.Unsafe.to_constr elimclause.templval.rebus)) with + (match EConstr.kind sigma (nth_arg sigma i elimclause.templval.rebus) with | Meta mv -> mv | _ -> user_err ~hdr:"elimination_clause" (str "The type of elimination clause is not well-formed.")) @@ -1421,7 +1445,7 @@ let elimination_clause_scheme with_evars ?(with_classes=true) ?(flags=elim_flags type eliminator = { elimindex : int option; (* None = find it automatically *) elimrename : (bool * int array) option; (** None = don't rename Prop hyps with H-names *) - elimbody : constr with_bindings + elimbody : EConstr.constr with_bindings } let general_elim_clause_gen elimtac indclause elim = @@ -1429,9 +1453,10 @@ let general_elim_clause_gen elimtac indclause elim = let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in let (elimc,lbindelimc) = elim.elimbody in - let elimt = Retyping.get_type_of env sigma (EConstr.of_constr elimc) in + let elimt = Retyping.get_type_of env sigma elimc in + let elimt = EConstr.of_constr elimt in let i = - match elim.elimindex with None -> index_of_ind_arg elimt | Some i -> i in + match elim.elimindex with None -> index_of_ind_arg sigma elimt | Some i -> i in elimtac elim.elimrename i (elimc, elimt, lbindelimc) indclause end } @@ -1439,12 +1464,11 @@ let general_elim with_evars clear_flag (c, lbindc) elim = Proofview.Goal.enter { enter = begin fun gl -> let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in - let ct = Retyping.get_type_of env sigma (EConstr.of_constr c) in - let t = try snd (reduce_to_quantified_ind env sigma (EConstr.of_constr ct)) with UserError _ -> ct in - let t = EConstr.of_constr t in + let ct = Retyping.get_type_of env sigma c in + let ct = EConstr.of_constr ct in + let t = try snd (reduce_to_quantified_ind env sigma ct) with UserError _ -> ct in let elimtac = elimination_clause_scheme with_evars in - let lbindc = Miscops.map_bindings EConstr.of_constr lbindc in - let indclause = make_clenv_binding env sigma (EConstr.of_constr c, t) lbindc in + let indclause = make_clenv_binding env sigma (c, t) lbindc in let sigma = meta_merge sigma (clear_metas indclause.evd) in Proofview.Unsafe.tclEVARS sigma <*> Tacticals.New.tclTHEN @@ -1459,15 +1483,16 @@ let general_case_analysis_in_context with_evars clear_flag (c,lbindc) = let sigma = Proofview.Goal.sigma gl in let env = Proofview.Goal.env gl in let concl = Proofview.Goal.concl gl in - let t = Retyping.get_type_of env (Sigma.to_evar_map sigma) (EConstr.of_constr c) in + let t = Retyping.get_type_of env (Sigma.to_evar_map sigma) c in let t = EConstr.of_constr t in let (mind,_) = reduce_to_quantified_ind env (Sigma.to_evar_map sigma) t in let sort = Tacticals.New.elimination_sort_of_goal gl in let Sigma (elim, sigma, p) = - if occur_term (Sigma.to_evar_map sigma) (EConstr.of_constr c) (EConstr.of_constr concl) then + if occur_term (Sigma.to_evar_map sigma) c (EConstr.of_constr concl) then build_case_analysis_scheme env sigma mind true sort else build_case_analysis_scheme_default env sigma mind sort in + let elim = EConstr.of_constr elim in let tac = (general_elim with_evars clear_flag (c,lbindc) {elimindex = None; elimbody = (elim,NoBindings); @@ -1477,7 +1502,8 @@ let general_case_analysis_in_context with_evars clear_flag (c,lbindc) = end } let general_case_analysis with_evars clear_flag (c,lbindc as cx) = - match kind_of_term c with + Proofview.tclEVARMAP >>= fun sigma -> + match EConstr.kind sigma c with | Var id when lbindc == NoBindings -> Tacticals.New.tclTHEN (try_intros_until_id_check id) (general_case_analysis_in_context with_evars clear_flag cx) @@ -1497,10 +1523,10 @@ let is_nonrec mind = (Global.lookup_mind (fst mind)).mind_finite == Decl_kinds.B let find_ind_eliminator ind s gl = let gr = lookup_eliminator ind s in let evd, c = Tacmach.New.pf_apply Evd.fresh_global gl gr in + let c = EConstr.of_constr c in evd, c let find_eliminator c gl = - let c = EConstr.of_constr c in let ((ind,u),t) = Tacmach.New.pf_reduce_to_quantified_ind gl (EConstr.of_constr (Tacmach.New.pf_unsafe_type_of gl c)) in if is_nonrec ind then raise IsNonrec; let evd, c = find_ind_eliminator ind (Tacticals.New.elimination_sort_of_goal gl) gl in @@ -1531,7 +1557,8 @@ let elim_in_context with_evars clear_flag c = function | None -> default_elim with_evars clear_flag c let elim with_evars clear_flag (c,lbindc as cx) elim = - match kind_of_term c with + Proofview.tclEVARMAP >>= fun sigma -> + match EConstr.kind sigma c with | Var id when lbindc == NoBindings -> Tacticals.New.tclTHEN (try_intros_until_id_check id) (elim_in_context with_evars clear_flag cx elim) @@ -1565,12 +1592,9 @@ let elimination_in_clause_scheme with_evars ?(flags=elim_flags ()) Proofview.Goal.enter { enter = begin fun gl -> let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in - let elim = contract_letin_in_lam_header elim in - let elim = EConstr.of_constr elim in - let elimty = EConstr.of_constr elimty in - let bindings = Miscops.map_bindings EConstr.of_constr bindings in + let elim = contract_letin_in_lam_header sigma elim in let elimclause = make_clenv_binding env sigma (elim, elimty) bindings in - let indmv = destMeta (nth_arg i (EConstr.Unsafe.to_constr elimclause.templval.rebus)) in + let indmv = destMeta sigma (nth_arg sigma i elimclause.templval.rebus) in let hypmv = try match List.remove Int.equal indmv (clenv_independent elimclause) with | [a] -> a @@ -1578,7 +1602,7 @@ let elimination_in_clause_scheme with_evars ?(flags=elim_flags ()) with Failure _ -> user_err ~hdr:"elimination_clause" (str "The type of elimination clause is not well-formed.") in let elimclause' = clenv_fchain ~flags indmv elimclause indclause in - let hyp = EConstr.mkVar id in + let hyp = mkVar id in let hyp_typ = Retyping.get_type_of env sigma hyp in let hyp_typ = EConstr.of_constr hyp_typ in let hypclause = mk_clenv_from_env env sigma (Some 0) (hyp, hyp_typ) in @@ -1611,19 +1635,23 @@ let make_projection env sigma params cstr sign elim i n c u = (* bugs: goes from right to left when i increases! *) let decl = List.nth cstr.cs_args i in let t = RelDecl.get_type decl in - let b = match decl with LocalAssum _ -> mkRel (i+1) | LocalDef (_,b,_) -> b in + let t = EConstr.of_constr t in + let b = match decl with LocalAssum _ -> mkRel (i+1) | LocalDef (_,b,_) -> EConstr.of_constr b in let branch = it_mkLambda_or_LetIn b cstr.cs_args in if (* excludes dependent projection types *) - noccur_between 1 (n-i-1) t + noccur_between sigma 1 (n-i-1) t (* to avoid surprising unifications, excludes flexible projection types or lambda which will be instantiated by Meta/Evar *) - && not (EConstr.isEvar sigma (fst (whd_betaiota_stack sigma (EConstr.of_constr t)))) - && (accept_universal_lemma_under_conjunctions () || not (isRel t)) + && not (isEvar sigma (fst (whd_betaiota_stack sigma t))) + && (accept_universal_lemma_under_conjunctions () || not (isRel sigma t)) then let t = lift (i+1-n) t in - let abselim = beta_applist sigma (EConstr.of_constr elim, List.map EConstr.of_constr (params@[t;branch])) in - let c = beta_applist sigma (EConstr.of_constr abselim, [EConstr.of_constr (mkApp (c, Context.Rel.to_extended_vect 0 sign))]) in + let abselim = beta_applist sigma (elim, params@[t;branch]) in + let abselim = EConstr.of_constr abselim in + let args = Array.map EConstr.of_constr (Context.Rel.to_extended_vect 0 sign) in + let c = beta_applist sigma (abselim, [mkApp (c, args)]) in + let c = EConstr.of_constr c in Some (it_mkLambda_or_LetIn c sign, it_mkProd_or_LetIn t sign) else None @@ -1632,6 +1660,7 @@ let make_projection env sigma params cstr sign elim i n c u = match List.nth l i with | Some proj -> let args = Context.Rel.to_extended_vect 0 sign in + let args = Array.map EConstr.of_constr args in let proj = if Environ.is_projection proj env then mkProj (Projection.make proj false, mkApp (c, args)) @@ -1640,7 +1669,8 @@ let make_projection env sigma params cstr sign elim i n c u = [|mkApp (c, args)|]) in let app = it_mkLambda_or_LetIn proj sign in - let t = Retyping.get_type_of env sigma (EConstr.of_constr app) in + let t = Retyping.get_type_of env sigma app in + let t = EConstr.of_constr t in Some (app, t) | None -> None in elim @@ -1650,23 +1680,24 @@ let descend_in_conjunctions avoid tac (err, info) c = let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in try - let t = Retyping.get_type_of env sigma (EConstr.of_constr c) in + let t = Retyping.get_type_of env sigma c in let t = EConstr.of_constr t in let ((ind,u),t) = reduce_to_quantified_ind env sigma t in - let sign,ccl = decompose_prod_assum t in - let ccl = EConstr.of_constr ccl in + let sign,ccl = EConstr.decompose_prod_assum sigma t in match match_with_tuple sigma ccl with | Some (_,_,isrec) -> let n = (constructors_nrealargs ind).(0) in let sort = Tacticals.New.elimination_sort_of_goal gl in let IndType (indf,_) = find_rectype env sigma ccl in let (_,inst), params = dest_ind_family indf in + let params = List.map EConstr.of_constr params in let cstr = (get_constructors env indf).(0) in let elim = try DefinedRecord (Recordops.lookup_projections ind) with Not_found -> let sigma = Sigma.Unsafe.of_evar_map sigma in let Sigma (elim, _, _) = build_case_analysis_scheme env sigma (ind,u) false sort in + let elim = EConstr.of_constr elim in NotADefinedRecordUseScheme elim in Tacticals.New.tclORELSE0 (Tacticals.New.tclFIRST @@ -1677,7 +1708,6 @@ let descend_in_conjunctions avoid tac (err, info) c = match make_projection env sigma params cstr sign elim i n c u with | None -> Tacticals.New.tclFAIL 0 (mt()) | Some (p,pt) -> - let p = EConstr.of_constr p in Tacticals.New.tclTHENS (assert_before_gen false (NamingAvoid avoid) pt) [Proofview.V82.tactic (refine p); @@ -1720,7 +1750,7 @@ let tclORELSEOPT t k = Proofview.tclZERO ~info e | Some tac -> tac) -let general_apply with_delta with_destruct with_evars clear_flag (loc,(c,lbind)) = +let general_apply with_delta with_destruct with_evars clear_flag (loc,(c,lbind : EConstr.constr with_bindings)) = Proofview.Goal.nf_enter { enter = begin fun gl -> let concl = Proofview.Goal.concl gl in let sigma = Tacmach.New.project gl in @@ -1735,14 +1765,13 @@ let general_apply with_delta with_destruct with_evars clear_flag (loc,(c,lbind)) let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in - let thm_ty0 = nf_betaiota sigma (EConstr.of_constr (Retyping.get_type_of env sigma (EConstr.of_constr c))) in + let thm_ty0 = nf_betaiota sigma (EConstr.of_constr (Retyping.get_type_of env sigma c)) in let try_apply thm_ty nprod = try let thm_ty = EConstr.of_constr thm_ty in let n = nb_prod_modulo_zeta sigma thm_ty - nprod in if n<0 then error "Applied theorem has not enough premisses."; - let lbind = Miscops.map_bindings EConstr.of_constr lbind in - let clause = make_clenv_binding_apply env sigma (Some n) (EConstr.of_constr c,thm_ty) lbind in + let clause = make_clenv_binding_apply env sigma (Some n) (c,thm_ty) lbind in Clenvtac.res_pf clause ~with_evars ~flags with exn when catchable_exception exn -> Proofview.tclZERO exn @@ -1863,7 +1892,6 @@ let progress_with_clause flags innerclause clause = with Not_found -> error "Unable to unify." let apply_in_once_main flags innerclause env sigma (d,lbind) = - let d = EConstr.of_constr d in let thm = nf_betaiota sigma (EConstr.of_constr (Retyping.get_type_of env sigma d)) in let thm = EConstr.of_constr thm in let rec aux clause = @@ -1873,7 +1901,6 @@ let apply_in_once_main flags innerclause env sigma (d,lbind) = try aux (clenv_push_prod clause) with NotExtensibleClause -> iraise e in - let lbind = Miscops.map_bindings EConstr.of_constr lbind in aux (make_clenv_binding env sigma (d,thm) lbind) let apply_in_once sidecond_first with_delta with_destruct with_evars naming @@ -1885,8 +1912,9 @@ let apply_in_once sidecond_first with_delta with_destruct with_evars naming let flags = if with_delta then default_unify_flags () else default_no_delta_unify_flags () in let t' = Tacmach.New.pf_get_hyp_typ id gl in - let innerclause = mk_clenv_from_env env sigma (Some 0) (EConstr.mkVar id,EConstr.of_constr t') in - let targetid = find_name true (LocalAssum (Anonymous,t')) naming gl in + let t' = EConstr.of_constr t' in + let innerclause = mk_clenv_from_env env sigma (Some 0) (mkVar id,t') in + let targetid = find_name true (local_assum (Anonymous,t')) naming gl in let rec aux idstoclear with_destruct c = Proofview.Goal.enter { enter = begin fun gl -> let env = Proofview.Goal.env gl in @@ -1942,16 +1970,16 @@ let apply_in_delayed_once sidecond_first with_delta with_destruct with_evars nam let cut_and_apply c = Proofview.Goal.nf_enter { enter = begin fun gl -> let sigma = Tacmach.New.project gl in - match kind_of_term (Tacmach.New.pf_hnf_constr gl (EConstr.of_constr (Tacmach.New.pf_unsafe_type_of gl (EConstr.of_constr c)))) with - | Prod (_,c1,c2) when EConstr.Vars.noccurn sigma 1 (EConstr.of_constr c2) -> + match EConstr.kind sigma (EConstr.of_constr (Tacmach.New.pf_hnf_constr gl (EConstr.of_constr (Tacmach.New.pf_unsafe_type_of gl c)))) with + | Prod (_,c1,c2) when Vars.noccurn sigma 1 c2 -> let concl = Proofview.Goal.concl gl in + let concl = EConstr.of_constr concl in let env = Tacmach.New.pf_env gl in Refine.refine { run = begin fun sigma -> let typ = mkProd (Anonymous, c2, concl) in - let Sigma (f, sigma, p) = Evarutil.new_evar env sigma (EConstr.of_constr typ) in - let Sigma (x, sigma, q) = Evarutil.new_evar env sigma (EConstr.of_constr c1) in + let Sigma (f, sigma, p) = Evarutil.new_evar env sigma typ in + let Sigma (x, sigma, q) = Evarutil.new_evar env sigma c1 in let ans = mkApp (f, [|mkApp (c, [|x|])|]) in - let ans = EConstr.of_constr ans in Sigma (ans, sigma, p +> q) end } | _ -> error "lapply needs a non-dependent product." @@ -1968,7 +1996,6 @@ let cut_and_apply c = (* let refine_no_check = Profile.profile2 refine_no_checkkey refine_no_check *) let exact_no_check c = - let c = EConstr.of_constr c in Refine.refine ~unsafe:true { run = fun h -> Sigma.here c h } let exact_check c = @@ -1976,9 +2003,11 @@ let exact_check c = let sigma = Proofview.Goal.sigma gl in (** We do not need to normalize the goal because we just check convertibility *) let concl = Proofview.Goal.concl (Proofview.Goal.assume gl) in + let concl = EConstr.of_constr concl in let env = Proofview.Goal.env gl in let sigma = Sigma.to_evar_map sigma in - let sigma, ct = Typing.type_of env sigma (EConstr.of_constr c) in + let sigma, ct = Typing.type_of env sigma c in + let ct = EConstr.of_constr ct in let tac = Tacticals.New.tclTHEN (convert_leq ct concl) (exact_no_check c) in @@ -1988,7 +2017,8 @@ let exact_check c = let cast_no_check cast c = Proofview.Goal.enter { enter = begin fun gl -> let concl = Proofview.Goal.concl (Proofview.Goal.assume gl) in - exact_no_check (Term.mkCast (c, cast, concl)) + let concl = EConstr.of_constr concl in + exact_no_check (EConstr.mkCast (c, cast, concl)) end } let vm_cast_no_check c = cast_no_check Term.VMcast c @@ -2048,7 +2078,7 @@ exception DependsOnBody of Id.t option let check_is_type env sigma ty = let evdref = ref sigma in try - let _ = Typing.e_sort_of env evdref (EConstr.of_constr ty) in + let _ = Typing.e_sort_of env evdref ty in !evdref with e when CErrors.noncritical e -> raise (DependsOnBody None) @@ -2073,6 +2103,7 @@ let clear_body ids = Proofview.Goal.enter { enter = begin fun gl -> let env = Proofview.Goal.env gl in let concl = Proofview.Goal.concl (Proofview.Goal.assume gl) in + let concl = EConstr.of_constr concl in let sigma = Tacmach.New.project gl in let ctx = named_context env in let map = function @@ -2102,7 +2133,7 @@ let clear_body ids = in let (env, sigma, _) = List.fold_left check (base_env, sigma, false) (List.rev ctx) in let sigma = - if List.exists (fun id -> occur_var env sigma id (EConstr.of_constr concl)) ids then + if List.exists (fun id -> occur_var env sigma id concl) ids then check_is_type env sigma concl else sigma in @@ -2116,8 +2147,7 @@ let clear_body ids = in check <*> Refine.refine ~unsafe:true { run = begin fun sigma -> - let Sigma (c, sigma, p) = Evarutil.new_evar env sigma ~principal:true (EConstr.of_constr concl) in - Sigma (EConstr.of_constr c, sigma, p) + Evarutil.new_evar env sigma ~principal:true concl end } end } @@ -2168,10 +2198,11 @@ let apply_type newcl args = let env = Proofview.Goal.env gl in let store = Proofview.Goal.extra gl in Refine.refine { run = begin fun sigma -> - let newcl = nf_betaiota (Sigma.to_evar_map sigma) (EConstr.of_constr newcl) (* As in former Logic.refine *) in + let newcl = nf_betaiota (Sigma.to_evar_map sigma) newcl (* As in former Logic.refine *) in + let newcl = EConstr.of_constr newcl in let Sigma (ev, sigma, p) = - Evarutil.new_evar env sigma ~principal:true ~store (EConstr.of_constr newcl) in - Sigma (EConstr.of_constr (applist (ev, args)), sigma, p) + Evarutil.new_evar env sigma ~principal:true ~store newcl in + Sigma (applist (ev, args), sigma, p) end } end } @@ -2186,12 +2217,13 @@ let bring_hyps hyps = let env = Proofview.Goal.env gl in let store = Proofview.Goal.extra gl in let concl = Tacmach.New.pf_nf_concl gl in + let concl = EConstr.of_constr concl in let newcl = List.fold_right mkNamedProd_or_LetIn hyps concl in - let args = Array.of_list (Context.Named.to_instance hyps) in + let args = Array.map_of_list EConstr.of_constr (Context.Named.to_instance hyps) in Refine.refine { run = begin fun sigma -> let Sigma (ev, sigma, p) = - Evarutil.new_evar env sigma ~principal:true ~store (EConstr.of_constr newcl) in - Sigma (EConstr.of_constr (mkApp (ev, args)), sigma, p) + Evarutil.new_evar env sigma ~principal:true ~store newcl in + Sigma (mkApp (ev, args), sigma, p) end } end } @@ -2322,10 +2354,10 @@ let my_find_eq_data_decompose gl t = let intro_decomp_eq loc l thin tac id = Proofview.Goal.nf_enter { enter = begin fun gl -> let c = mkVar id in - let t = Tacmach.New.pf_unsafe_type_of gl (EConstr.of_constr c) in + let t = Tacmach.New.pf_unsafe_type_of gl c in let t = EConstr.of_constr t in let _,t = Tacmach.New.pf_reduce_to_quantified_ind gl t in - match my_find_eq_data_decompose gl (EConstr.of_constr t) with + match my_find_eq_data_decompose gl t with | Some (eq,u,eq_args) -> !intro_decomp_eq_function (fun n -> tac ((dloc,id)::thin) (Some (true,n)) l) @@ -2337,7 +2369,7 @@ let intro_decomp_eq loc l thin tac id = let intro_or_and_pattern loc with_evars bracketed ll thin tac id = Proofview.Goal.enter { enter = begin fun gl -> let c = mkVar id in - let t = Tacmach.New.pf_unsafe_type_of gl (EConstr.of_constr c) in + let t = Tacmach.New.pf_unsafe_type_of gl c in let t = EConstr.of_constr t in let (ind,t) = Tacmach.New.pf_reduce_to_quantified_ind gl t in let branchsigns = compute_constructor_signatures false ind in @@ -2363,26 +2395,23 @@ let rewrite_hyp_then assert_style with_evars thin l2r id tac = let sigma = Tacmach.New.project gl in let type_of = Tacmach.New.pf_unsafe_type_of gl in let whd_all = Tacmach.New.pf_apply whd_all gl in - let t = whd_all (EConstr.of_constr (type_of (EConstr.mkVar id))) in + let t = whd_all (EConstr.of_constr (type_of (mkVar id))) in let t = EConstr.of_constr t in let eqtac, thin = match match_with_equality_type sigma t with | Some (hdcncl,[_;lhs;rhs]) -> - let lhs = EConstr.Unsafe.to_constr lhs in - let rhs = EConstr.Unsafe.to_constr rhs in - if l2r && isVar lhs && not (occur_var env sigma (destVar lhs) (EConstr.of_constr rhs)) then - let id' = destVar lhs in + if l2r && isVar sigma lhs && not (occur_var env sigma (destVar sigma lhs) rhs) then + let id' = destVar sigma lhs in subst_on l2r id' rhs, early_clear id' thin - else if not l2r && isVar rhs && not (occur_var env sigma (destVar rhs) (EConstr.of_constr lhs)) then - let id' = destVar rhs in + else if not l2r && isVar sigma rhs && not (occur_var env sigma (destVar sigma rhs) lhs) then + let id' = destVar sigma rhs in subst_on l2r id' lhs, early_clear id' thin else Tacticals.New.tclTHEN (rew_on l2r onConcl) (clear [id]), thin | Some (hdcncl,[c]) -> - let c = EConstr.Unsafe.to_constr c in let l2r = not l2r in (* equality of the form eq_true *) - if isVar c then - let id' = destVar c in + if isVar sigma c then + let id' = destVar sigma c in Tacticals.New.tclTHEN (rew_on l2r allHypsAndConcl) (clear_var_and_eq id'), early_clear id' thin @@ -2581,9 +2610,9 @@ let ipat_of_name = function | Anonymous -> None | Name id -> Some (dloc, IntroNaming (IntroIdentifier id)) -let head_ident c = - let c = fst (decompose_app ((strip_lam_assum c))) in - if isVar c then Some (destVar c) else None +let head_ident sigma c = + let c = fst (decompose_app sigma (snd (decompose_lam_assum sigma c))) in + if isVar sigma c then Some (destVar sigma c) else None let assert_as first hd ipat t = let naming,tac = prepare_intros false IntroAnonymous MoveLast ipat in @@ -2652,8 +2681,10 @@ let letin_tac_gen with_eq (id,depdecls,lastlhyp,ccl,c) ty = let Sigma (t, sigma, p) = match ty with | Some t -> Sigma.here t sigma | None -> - let t = EConstr.of_constr (typ_of env sigma (EConstr.of_constr c)) in + let t = typ_of env sigma c in + let t = EConstr.of_constr t in let sigma, c = Evarsolve.refresh_universes ~onlyalg:true (Some false) env (Sigma.to_evar_map sigma) t in + let c = EConstr.of_constr c in Sigma.Unsafe.of_pair (c, sigma) in let Sigma ((newcl, eq_tac), sigma, q) = match with_eq with @@ -2665,12 +2696,14 @@ let letin_tac_gen with_eq (id,depdecls,lastlhyp,ccl,c) ty = let eqdata = build_coq_eq_data () in let args = if lr then [t;mkVar id;c] else [t;c;mkVar id]in let Sigma (eq, sigma, p) = Sigma.fresh_global env sigma eqdata.eq in + let eq = EConstr.of_constr eq in let Sigma (refl, sigma, q) = Sigma.fresh_global env sigma eqdata.refl in + let refl = EConstr.of_constr refl in let eq = applist (eq,args) in let refl = applist (refl, [t;mkVar id]) in let term = mkNamedLetIn id c t (mkLetIn (Name heq, refl, eq, ccl)) in let sigma = Sigma.to_evar_map sigma in - let sigma, _ = Typing.type_of env sigma (EConstr.of_constr term) in + let sigma, _ = Typing.type_of env sigma term in let ans = term, Tacticals.New.tclTHEN (intro_gen (NamingMustBe (loc,heq)) (decode_hyp lastlhyp) true false) @@ -2704,9 +2737,9 @@ let insert_before decls lasthyp env = let mkletin_goal env sigma store with_eq dep (id,lastlhyp,ccl,c) ty = let open Context.Named.Declaration in - let t = match ty with Some t -> t | _ -> typ_of env sigma (EConstr.of_constr c) in - let decl = if dep then LocalDef (id,c,t) - else LocalAssum (id,t) + let t = match ty with Some t -> t | _ -> EConstr.of_constr (typ_of env sigma c) in + let decl = if dep then nlocal_def (id,c,t) + else nlocal_assum (id,t) in match with_eq with | Some (lr,(loc,ido)) -> @@ -2720,34 +2753,33 @@ let mkletin_goal env sigma store with_eq dep (id,lastlhyp,ccl,c) ty = let eqdata = build_coq_eq_data () in let args = if lr then [t;mkVar id;c] else [t;c;mkVar id]in let Sigma (eq, sigma, p) = Sigma.fresh_global env sigma eqdata.eq in + let eq = EConstr.of_constr eq in let Sigma (refl, sigma, q) = Sigma.fresh_global env sigma eqdata.refl in + let refl = EConstr.of_constr refl in let eq = applist (eq,args) in let refl = applist (refl, [t;mkVar id]) in - let newenv = insert_before [LocalAssum (heq,eq); decl] lastlhyp env in - let Sigma (x, sigma, r) = new_evar newenv sigma ~principal:true ~store (EConstr.of_constr ccl) in - Sigma (EConstr.of_constr (mkNamedLetIn id c t (mkNamedLetIn heq refl eq x)), sigma, p +> q +> r) + let newenv = insert_before [nlocal_assum (heq,eq); decl] lastlhyp env in + let Sigma (x, sigma, r) = new_evar newenv sigma ~principal:true ~store ccl in + Sigma (mkNamedLetIn id c t (mkNamedLetIn heq refl eq x), sigma, p +> q +> r) | None -> let newenv = insert_before [decl] lastlhyp env in - let Sigma (x, sigma, p) = new_evar newenv sigma ~principal:true ~store (EConstr.of_constr ccl) in - Sigma (EConstr.of_constr (mkNamedLetIn id c t x), sigma, p) + let Sigma (x, sigma, p) = new_evar newenv sigma ~principal:true ~store ccl in + Sigma (mkNamedLetIn id c t x, sigma, p) let letin_tac with_eq id c ty occs = Proofview.Goal.nf_s_enter { s_enter = begin fun gl -> let sigma = Proofview.Goal.sigma gl in let env = Proofview.Goal.env gl in let ccl = Proofview.Goal.concl gl in - let c = EConstr.of_constr c in - let abs = AbstractExact (id,c,Option.map EConstr.of_constr ty,occs,true) in + let abs = AbstractExact (id,c,ty,occs,true) in let ccl = EConstr.of_constr ccl in let (id,_,depdecls,lastlhyp,ccl,res) = make_abstraction env sigma ccl abs in - let ccl = EConstr.Unsafe.to_constr ccl in (* We keep the original term to match but record the potential side-effects of unifying universes. *) let Sigma (c, sigma, p) = match res with | None -> Sigma.here c sigma | Some (Sigma (_, sigma, p)) -> Sigma (c, sigma, p) in - let c = EConstr.Unsafe.to_constr c in let tac = letin_tac_gen with_eq (id,depdecls,lastlhyp,ccl,c) ty in Sigma (tac, sigma, p) end } @@ -2761,11 +2793,9 @@ let letin_pat_tac with_eq id c occs = let abs = AbstractPattern (false,check,id,c,occs,false) in let ccl = EConstr.of_constr ccl in let (id,_,depdecls,lastlhyp,ccl,res) = make_abstraction env sigma ccl abs in - let ccl = EConstr.Unsafe.to_constr ccl in let Sigma (c, sigma, p) = match res with | None -> finish_evar_resolution ~flags:(tactic_infer_flags false) env sigma c | Some res -> res in - let c = EConstr.Unsafe.to_constr c in let tac = (letin_tac_gen with_eq (id,depdecls,lastlhyp,ccl,c) None) in @@ -2777,8 +2807,10 @@ let forward b usetac ipat c = match usetac with | None -> Proofview.Goal.enter { enter = begin fun gl -> - let t = Tacmach.New.pf_unsafe_type_of gl (EConstr.of_constr c) in - let hd = head_ident c in + let t = Tacmach.New.pf_unsafe_type_of gl c in + let t = EConstr.of_constr t in + let sigma = Tacmach.New.project gl in + let hd = head_ident sigma c in Tacticals.New.tclTHENFIRST (assert_as true hd ipat t) (exact_no_check c) end } | Some tac -> @@ -2801,22 +2833,22 @@ let enough_by na t tac = forward false (Some (Some tac)) (ipat_of_name na) t (* Compute a name for a generalization *) -let generalized_name c t ids cl = function +let generalized_name sigma c t ids cl = function | Name id as na -> if Id.List.mem id ids then user_err (pr_id id ++ str " is already used."); na | Anonymous -> - match kind_of_term c with + match EConstr.kind sigma c with | Var id -> (* Keep the name even if not occurring: may be used by intros later *) Name id | _ -> - if noccurn 1 cl then Anonymous else + if noccurn sigma 1 cl then Anonymous else (* On ne s'etait pas casse la tete : on avait pris pour nom de variable la premiere lettre du type, meme si "c" avait ete une constante dont on aurait pu prendre directement le nom *) - named_hd (Global.env()) t Anonymous + named_hd (Global.env()) (EConstr.Unsafe.to_constr t) Anonymous (* Abstract over [c] in [forall x1:A1(c)..xi:Ai(c).T(c)] producing [forall x, x1:A1(x1), .., xi:Ai(x). T(x)] with all [c] abtracted in [Ai] @@ -2824,21 +2856,23 @@ let generalized_name c t ids cl = function let generalize_goal_gen env sigma ids i ((occs,c,b),na) t cl = let open Context.Rel.Declaration in - let decls,cl = decompose_prod_n_assum i cl in - let dummy_prod = EConstr.of_constr (it_mkProd_or_LetIn mkProp decls) in - let newdecls,_ = decompose_prod_n_assum i (subst_term_gen sigma EConstr.eq_constr_nounivs (EConstr.of_constr c) dummy_prod) in - let cl',sigma' = subst_closed_term_occ env sigma (AtOccs occs) (EConstr.of_constr c) (EConstr.of_constr (it_mkProd_or_LetIn cl newdecls)) in - let na = generalized_name c t ids cl' na in + let decls,cl = decompose_prod_n_assum sigma i cl in + let dummy_prod = it_mkProd_or_LetIn mkProp decls in + let newdecls,_ = decompose_prod_n_assum sigma i (EConstr.of_constr (subst_term_gen sigma EConstr.eq_constr_nounivs c dummy_prod)) in + let cl',sigma' = subst_closed_term_occ env sigma (AtOccs occs) c (it_mkProd_or_LetIn cl newdecls) in + let cl' = EConstr.of_constr cl' in + let na = generalized_name sigma c t ids cl' na in let decl = match b with - | None -> LocalAssum (na,t) - | Some b -> LocalDef (na,b,t) + | None -> local_assum (na,t) + | Some b -> local_def (na,b,t) in mkProd_or_LetIn decl cl', sigma' let generalize_goal gl i ((occs,c,b),na as o) (cl,sigma) = let env = Tacmach.pf_env gl in let ids = Tacmach.pf_ids_of_hyps gl in - let sigma, t = Typing.type_of env sigma (EConstr.of_constr c) in + let sigma, t = Typing.type_of env sigma c in + let t = EConstr.of_constr t in generalize_goal_gen env sigma ids i o t cl let old_generalize_dep ?(with_let=false) c gl = @@ -2848,7 +2882,7 @@ let old_generalize_dep ?(with_let=false) c gl = let init_ids = ids_of_named_context (Global.named_context()) in let seek (d:Context.Named.Declaration.t) (toquant:Context.Named.t) = if List.exists (fun d' -> occur_var_in_decl env sigma (NamedDecl.get_id d') d) toquant - || dependent_in_decl sigma (EConstr.of_constr c) d then + || dependent_in_decl sigma c d then d::toquant else toquant in @@ -2857,24 +2891,27 @@ let old_generalize_dep ?(with_let=false) c gl = let qhyps = List.map NamedDecl.get_id to_quantify_rev in let tothin = List.filter (fun id -> not (Id.List.mem id init_ids)) qhyps in let tothin' = - match kind_of_term c with + match EConstr.kind sigma c with | Var id when mem_named_context_val id (val_of_named_context sign) && not (Id.List.mem id init_ids) -> id::tothin | _ -> tothin in let cl' = it_mkNamedProd_or_LetIn (Tacmach.pf_concl gl) to_quantify in + let cl' = EConstr.of_constr cl' in let body = if with_let then - match kind_of_term c with + match EConstr.kind sigma c with | Var id -> id |> Tacmach.pf_get_hyp gl |> NamedDecl.get_value | _ -> None else None in + let body = Option.map EConstr.of_constr body in let cl'',evd = generalize_goal gl 0 ((AllOccurrences,c,body),Anonymous) (cl',project gl) in (** Check that the generalization is indeed well-typed *) - let (evd, _) = Typing.type_of env evd (EConstr.of_constr cl'') in + let (evd, _) = Typing.type_of env evd cl'' in let args = Context.Named.to_instance to_quantify_rev in + let args = List.map EConstr.of_constr args in tclTHENLIST [tclEVARS evd; Proofview.V82.of_tactic (apply_type cl'' (if Option.is_empty body then c::args else args)); @@ -2889,9 +2926,9 @@ let generalize_gen_let lconstr = Proofview.Goal.nf_s_enter { s_enter = begin fun let env = Proofview.Goal.env gl in let newcl, evd = List.fold_right_i (Tacmach.New.of_old generalize_goal gl) 0 lconstr - (Tacmach.New.pf_concl gl,Tacmach.New.project gl) + (EConstr.of_constr (Tacmach.New.pf_concl gl),Tacmach.New.project gl) in - let (evd, _) = Typing.type_of env evd (EConstr.of_constr newcl) in + let (evd, _) = Typing.type_of env evd newcl in let map ((_, c, b),_) = if Option.is_empty b then Some c else None in let tac = apply_type newcl (List.map_filter map lconstr) in Sigma.Unsafe.of_pair (tac, evd) @@ -2902,13 +2939,15 @@ let new_generalize_gen_let lconstr = let sigma = Proofview.Goal.sigma gl in let gl = Proofview.Goal.assume gl in let concl = Proofview.Goal.concl gl in + let concl = EConstr.of_constr concl in let sigma = Sigma.to_evar_map sigma in let env = Proofview.Goal.env gl in let ids = Tacmach.New.pf_ids_of_hyps gl in let newcl, sigma, args = List.fold_right_i (fun i ((_,c,b),_ as o) (cl, sigma, args) -> - let sigma, t = Typing.type_of env sigma (EConstr.of_constr c) in + let sigma, t = Typing.type_of env sigma c in + let t = EConstr.of_constr t in let args = if Option.is_empty b then c :: args else args in let cl, sigma = generalize_goal_gen env sigma ids i o t cl in (cl, sigma, args)) @@ -2916,8 +2955,8 @@ let new_generalize_gen_let lconstr = in let tac = Refine.refine { run = begin fun sigma -> - let Sigma (ev, sigma, p) = Evarutil.new_evar env sigma ~principal:true (EConstr.of_constr newcl) in - Sigma (EConstr.of_constr (applist (ev, args)), sigma, p) + let Sigma (ev, sigma, p) = Evarutil.new_evar env sigma ~principal:true newcl in + Sigma ((applist (ev, args)), sigma, p) end } in Sigma.Unsafe.of_pair (tac, sigma) @@ -2950,6 +2989,7 @@ let quantify lconstr = (* Modifying/Adding an hypothesis *) let specialize (c,lbind) ipat = + let nf_evar sigma c = EConstr.of_constr (nf_evar sigma (EConstr.Unsafe.to_constr c)) in Proofview.Goal.enter { enter = begin fun gl -> let env = Proofview.Goal.env gl in let sigma = Sigma.to_evar_map (Proofview.Goal.sigma gl) in @@ -2958,27 +2998,26 @@ let specialize (c,lbind) ipat = let sigma = Typeclasses.resolve_typeclasses env sigma in sigma, nf_evar sigma c else - let c = EConstr.of_constr c in - let lbind = Miscops.map_bindings EConstr.of_constr lbind in let clause = make_clenv_binding env sigma (c,EConstr.of_constr (Retyping.get_type_of env sigma c)) lbind in let flags = { (default_unify_flags ()) with resolve_evars = true } in let clause = clenv_unify_meta_types ~flags clause in let (thd,tstack) = whd_nored_stack clause.evd (clenv_value clause) in let rec chk = function | [] -> [] - | t::l -> if occur_meta clause.evd t then [] else EConstr.Unsafe.to_constr t :: chk l + | t::l -> if occur_meta clause.evd t then [] else t :: chk l in let tstack = chk tstack in - let term = applist(EConstr.Unsafe.to_constr thd,List.map (nf_evar clause.evd) tstack) in - if occur_meta clause.evd (EConstr.of_constr term) then + let term = applist(thd,List.map (nf_evar clause.evd) tstack) in + if occur_meta clause.evd term then user_err (str "Cannot infer an instance for " ++ - pr_name (meta_name clause.evd (List.hd (collect_metas clause.evd (EConstr.of_constr term)))) ++ + pr_name (meta_name clause.evd (List.hd (collect_metas clause.evd term))) ++ str "."); clause.evd, term in - let typ = Retyping.get_type_of env sigma (EConstr.of_constr term) in + let typ = Retyping.get_type_of env sigma term in + let typ = EConstr.of_constr typ in let tac = - match kind_of_term (fst(decompose_app (snd(decompose_lam_assum c)))) with + match EConstr.kind sigma (fst(EConstr.decompose_app sigma (snd(EConstr.decompose_lam_assum sigma c)))) with | Var id when Id.List.mem id (Tacmach.New.pf_ids_of_hyps gl) -> (* Like assert (id:=id args) but with the concept of specialization *) let naming,tac = @@ -3020,9 +3059,10 @@ let unfold_body x = (pr_id x ++ str" is not a defined hypothesis.") | LocalDef (_,xval,_) -> xval in + let xval = EConstr.of_constr xval in Tacticals.New.afterHyp x begin fun aft -> let hl = List.fold_right (fun decl cl -> (NamedDecl.get_id decl, InHyp) :: cl) aft [] in - let rfun _ _ c = replace_vars [x, xval] (EConstr.Unsafe.to_constr c) in + let rfun _ _ c = EConstr.Unsafe.to_constr (replace_vars [x, xval] c) in let reducth h = reduct_in_hyp rfun h in let reductc = reduct_in_concl (rfun, DEFAULTcast) in Tacticals.New.tclTHENLIST [Tacticals.New.tclMAP reducth hl; reductc] @@ -3072,7 +3112,7 @@ let warn_unused_intro_pattern = strbrk"Unused introduction " ++ str (String.plural (List.length names) "pattern") ++ str": " ++ prlist_with_sep spc (Miscprint.pr_intro_pattern - (fun c -> Printer.pr_constr (fst (run_delayed (Global.env()) Evd.empty c)))) names) + (fun c -> Printer.pr_constr (EConstr.Unsafe.to_constr (fst (run_delayed (Global.env()) Evd.empty c))))) names) let check_unused_names names = if not (List.is_empty names) && Flags.is_verbose () then @@ -3206,13 +3246,12 @@ let induct_discharge with_evars dests avoid' tac (avoid,ra) names = substitutions aussi sur l'argument voisin *) let expand_projections env sigma c = - let sigma = Sigma.to_evar_map sigma in let rec aux env c = match EConstr.kind sigma c with | Proj (p, c) -> Retyping.expand_projection env sigma p (aux env c) [] | _ -> map_constr_with_full_binders sigma push_rel aux env c in - EConstr.Unsafe.to_constr (aux env (EConstr.of_constr c)) + aux env c (* Marche pas... faut prendre en compte l'occurrence précise... *) @@ -3220,13 +3259,14 @@ let expand_projections env sigma c = let atomize_param_of_ind_then (indref,nparams,_) hyp0 tac = Proofview.Goal.enter { enter = begin fun gl -> let env = Proofview.Goal.env gl in + let sigma = Tacmach.New.project gl in let tmptyp0 = Tacmach.New.pf_get_hyp_typ hyp0 (Proofview.Goal.assume gl) in + let tmptyp0 = EConstr.of_constr tmptyp0 in let reduce_to_quantified_ref = Tacmach.New.pf_apply reduce_to_quantified_ref gl in - let typ0 = reduce_to_quantified_ref indref (EConstr.of_constr tmptyp0) in - let prods, indtyp = decompose_prod_assum typ0 in - let hd,argl = decompose_app indtyp in + let typ0 = reduce_to_quantified_ref indref tmptyp0 in + let prods, indtyp = decompose_prod_assum sigma typ0 in + let hd,argl = decompose_app sigma indtyp in let env' = push_rel_context prods env in - let sigma = Proofview.Goal.sigma gl in let params = List.firstn nparams argl in let params' = List.map (expand_projections env' sigma) params in (* le gl est important pour ne pas préévaluer *) @@ -3238,16 +3278,16 @@ let atomize_param_of_ind_then (indref,nparams,_) hyp0 tac = (tac avoid) else let c = List.nth argl (i-1) in - match kind_of_term c with - | Var id when not (List.exists (fun c -> occur_var env (Sigma.to_evar_map sigma) id (EConstr.of_constr c)) args') && - not (List.exists (fun c -> occur_var env (Sigma.to_evar_map sigma) id (EConstr.of_constr c)) params') -> + match EConstr.kind sigma c with + | Var id when not (List.exists (fun c -> occur_var env sigma id c) args') && + not (List.exists (fun c -> occur_var env sigma id c) params') -> (* Based on the knowledge given by the user, all constraints on the variable are generalizable in the current environment so that it is clearable after destruction *) atomize_one (i-1) (c::args) (c::args') (id::avoid) | _ -> let c' = expand_projections env' sigma c in - let dependent t = dependent (Sigma.to_evar_map sigma) (EConstr.of_constr c) (EConstr.of_constr t) in + let dependent t = dependent sigma c t in if List.exists dependent params' || List.exists dependent args' then @@ -3261,11 +3301,11 @@ let atomize_param_of_ind_then (indref,nparams,_) hyp0 tac = (* We reason blindly on the term and do as if it were generalizable, ignoring the constraints coming from its structure *) - let id = match kind_of_term c with + let id = match EConstr.kind sigma c with | Var id -> id | _ -> let type_of = Tacmach.New.pf_unsafe_type_of gl in - id_of_name_using_hdchar (Global.env()) (type_of (EConstr.of_constr c)) Anonymous in + id_of_name_using_hdchar (Global.env()) (type_of c) Anonymous in let x = fresh_id_in_env avoid id env in Tacticals.New.tclTHEN (letin_tac None (Name x) c None allHypsAndConcl) @@ -3440,8 +3480,8 @@ let cook_sign hyp0_opt inhyps indvars env sigma = (* [rel_contexts] and [rel_declaration] actually contain triples, and lists are actually in reverse order to fit [compose_prod]. *) type elim_scheme = { - elimc: constr with_bindings option; - elimt: types; + elimc: EConstr.constr with_bindings option; + elimt: EConstr.types; indref: global_reference option; params: Context.Rel.t; (* (prm1,tprm1);(prm2,tprm2)...(prmp,tprmp) *) nparams: int; (* number of parameters *) @@ -3453,7 +3493,7 @@ type elim_scheme = { nargs: int; (* number of arguments *) indarg: Context.Rel.Declaration.t option; (* Some (H,I prm1..prmp x1...xni) if HI is in premisses, None otherwise *) - concl: types; (* Qi x1...xni HI (f...), HI and (f...) + concl: EConstr.types; (* Qi x1...xni HI (f...), HI and (f...) are optional and mutually exclusive *) indarg_in_concl: bool; (* true if HI appears at the end of conclusion *) farg_in_concl: bool; (* true if (f...) appears at the end of conclusion *) @@ -3462,7 +3502,7 @@ type elim_scheme = { let empty_scheme = { elimc = None; - elimt = mkProp; + elimt = EConstr.mkProp; indref = None; params = []; nparams = 0; @@ -3473,7 +3513,7 @@ let empty_scheme = args = []; nargs = 0; indarg = None; - concl = mkProp; + concl = EConstr.mkProp; indarg_in_concl = false; farg_in_concl = false; } @@ -3516,13 +3556,13 @@ let error_ind_scheme s = let s = if not (String.is_empty s) then s^" " else s in user_err ~hdr:"Tactics" (str "Cannot recognize " ++ str s ++ str "an induction scheme.") -let glob = Universes.constr_of_global +let glob c = EConstr.of_constr (Universes.constr_of_global c) let coq_eq = lazy (glob (Coqlib.build_coq_eq ())) let coq_eq_refl = lazy (glob (Coqlib.build_coq_eq_refl ())) -let coq_heq = lazy (Coqlib.coq_constant "mkHEq" ["Logic";"JMeq"] "JMeq") -let coq_heq_refl = lazy (Coqlib.coq_constant "mkHEq" ["Logic";"JMeq"] "JMeq_refl") +let coq_heq = lazy (EConstr.of_constr (Coqlib.coq_constant "mkHEq" ["Logic";"JMeq"] "JMeq")) +let coq_heq_refl = lazy (EConstr.of_constr (Coqlib.coq_constant "mkHEq" ["Logic";"JMeq"] "JMeq_refl")) let mkEq t x y = @@ -3547,26 +3587,26 @@ let lift_togethern n l = l ([], n) in l' -let lift_list l = List.map (lift 1) l +let lift_list l = List.map (EConstr.Vars.lift 1) l -let ids_of_constr ?(all=false) vars c = +let ids_of_constr sigma ?(all=false) vars c = let rec aux vars c = - match kind_of_term c with + match EConstr.kind sigma c with | Var id -> Id.Set.add id vars | App (f, args) -> - (match kind_of_term f with + (match EConstr.kind sigma f with | Construct ((ind,_),_) | Ind (ind,_) -> let (mib,mip) = Global.lookup_inductive ind in Array.fold_left_from (if all then 0 else mib.Declarations.mind_nparams) aux vars args - | _ -> Term.fold_constr aux vars c) - | _ -> Term.fold_constr aux vars c + | _ -> EConstr.fold sigma aux vars c) + | _ -> EConstr.fold sigma aux vars c in aux vars c -let decompose_indapp f args = - match kind_of_term f with +let decompose_indapp sigma f args = + match EConstr.kind sigma f with | Construct ((ind,_),_) | Ind (ind,_) -> let (mib,mip) = Global.lookup_inductive ind in @@ -3577,7 +3617,7 @@ let decompose_indapp f args = let mk_term_eq env sigma ty t ty' t' = let sigma = Sigma.to_evar_map sigma in - if Reductionops.is_conv env sigma (EConstr.of_constr ty) (EConstr.of_constr ty') then + if Reductionops.is_conv env sigma ty ty' then mkEq ty t t', mkRefl ty' t' else mkHEq ty t ty' t', mkHRefl ty' t' @@ -3595,17 +3635,17 @@ let make_abstract_generalize env id typ concl dep ctx body c eqs args refls = in (* Abstract by equalities *) let eqs = lift_togethern 1 eqs in (* lift together and past genarg *) - let abseqs = it_mkProd_or_LetIn (lift eqslen abshypeq) (List.map (fun x -> LocalAssum (Anonymous, x)) eqs) in + let abseqs = it_mkProd_or_LetIn (lift eqslen abshypeq) (List.map (fun x -> local_assum (Anonymous, x)) eqs) in let decl = match body with - | None -> LocalAssum (Name id, c) - | Some body -> LocalDef (Name id, body, c) + | None -> local_assum (Name id, c) + | Some body -> local_def (Name id, body, c) in (* Abstract by the "generalized" hypothesis. *) let genarg = mkProd_or_LetIn decl abseqs in (* Abstract by the extension of the context *) let genctyp = it_mkProd_or_LetIn genarg ctx in (* The goal will become this product. *) - let Sigma (genc, sigma, p) = Evarutil.new_evar env sigma ~principal:true (EConstr.of_constr genctyp) in + let Sigma (genc, sigma, p) = Evarutil.new_evar env sigma ~principal:true genctyp in (* Apply the old arguments giving the proper instantiation of the hyp *) let instc = mkApp (genc, Array.of_list args) in (* Then apply to the original instantiated hyp. *) @@ -3613,7 +3653,7 @@ let make_abstract_generalize env id typ concl dep ctx body c eqs args refls = (* Apply the reflexivity proofs on the indices. *) let appeqs = mkApp (instc, Array.of_list refls) in (* Finally, apply the reflexivity proof for the original hyp, to get a term of type gl again. *) - Sigma (EConstr.of_constr (mkApp (appeqs, abshypt)), sigma, p) + Sigma (mkApp (appeqs, abshypt), sigma, p) end } let hyps_of_vars env sigma sign nogen hyps = @@ -3636,11 +3676,11 @@ let hyps_of_vars env sigma sign nogen hyps = exception Seen -let linear vars args = +let linear sigma vars args = let seen = ref vars in try Array.iter (fun i -> - let rels = ids_of_constr ~all:true Id.Set.empty i in + let rels = ids_of_constr sigma ~all:true Id.Set.empty i in let seen' = Id.Set.fold (fun id acc -> if Id.Set.mem id acc then raise Seen @@ -3659,7 +3699,8 @@ let abstract_args gl generalize_vars dep id defined f args = let sigma = ref (Tacmach.project gl) in let env = Tacmach.pf_env gl in let concl = Tacmach.pf_concl gl in - let dep = dep || local_occur_var !sigma id (EConstr.of_constr concl) in + let concl = EConstr.of_constr concl in + let dep = dep || local_occur_var !sigma id concl in let avoid = ref [] in let get_id name = let id = fresh_id !avoid (match name with Name n -> n | Anonymous -> Id.of_string "gen_x") gl in @@ -3673,23 +3714,27 @@ let abstract_args gl generalize_vars dep id defined f args = *) let aux (prod, ctx, ctxenv, c, args, eqs, refls, nongenvars, vars, env) arg = let name, ty, arity = - let rel, c = Reductionops.splay_prod_n env !sigma 1 (EConstr.of_constr prod) in + let rel, c = Reductionops.splay_prod_n env !sigma 1 prod in + let c = EConstr.of_constr c in let decl = List.hd rel in RelDecl.get_name decl, RelDecl.get_type decl, c in - let argty = Tacmach.pf_unsafe_type_of gl (EConstr.of_constr arg) in - let sigma', ty = Evarsolve.refresh_universes (Some true) env !sigma (EConstr.of_constr ty) in + let ty = EConstr.of_constr ty in + let argty = Tacmach.pf_unsafe_type_of gl arg in + let argty = EConstr.of_constr argty in + let sigma', ty = Evarsolve.refresh_universes (Some true) env !sigma ty in let () = sigma := sigma' in + let ty = EConstr.of_constr ty in let lenctx = List.length ctx in let liftargty = lift lenctx argty in - let leq = constr_cmp Reduction.CUMUL liftargty ty in - match kind_of_term arg with + let leq = constr_cmp !sigma Reduction.CUMUL liftargty ty in + match EConstr.kind !sigma arg with | Var id when not (is_defined_variable env id) && leq && not (Id.Set.mem id nongenvars) -> (subst1 arg arity, ctx, ctxenv, mkApp (c, [|arg|]), args, eqs, refls, Id.Set.add id nongenvars, Id.Set.remove id vars, env) | _ -> let name = get_id name in - let decl = LocalAssum (Name name, ty) in + let decl = local_assum (Name name, ty) in let ctx = decl :: ctx in let c' = mkApp (lift 1 c, [|mkRel 1|]) in let args = arg :: args in @@ -3702,23 +3747,24 @@ let abstract_args gl generalize_vars dep id defined f args = in let eqs = eq :: lift_list eqs in let refls = refl :: refls in - let argvars = ids_of_constr vars arg in + let argvars = ids_of_constr !sigma vars arg in (arity, ctx, push_rel decl ctxenv, c', args, eqs, refls, nongenvars, Id.Set.union argvars vars, env) in - let f', args' = decompose_indapp f args in + let f', args' = decompose_indapp !sigma f args in let dogen, f', args' = - let parvars = ids_of_constr ~all:true Id.Set.empty f' in - if not (linear parvars args') then true, f, args + let parvars = ids_of_constr !sigma ~all:true Id.Set.empty f' in + if not (linear !sigma parvars args') then true, f, args else - match Array.findi (fun i x -> not (isVar x) || is_defined_variable env (destVar x)) args' with + match Array.findi (fun i x -> not (isVar !sigma x) || is_defined_variable env (destVar !sigma x)) args' with | None -> false, f', args' | Some nonvar -> let before, after = Array.chop nonvar args' in true, mkApp (f', before), after in if dogen then - let tyf' = Tacmach.pf_unsafe_type_of gl (EConstr.of_constr f') in + let tyf' = Tacmach.pf_unsafe_type_of gl f' in + let tyf' = EConstr.of_constr tyf' in let arity, ctx, ctxenv, c', args, eqs, refls, nogen, vars, env = Array.fold_left aux (tyf',[],env,f',[],[],[],Id.Set.empty,Id.Set.empty,env) args' in @@ -3730,10 +3776,11 @@ let abstract_args gl generalize_vars dep id defined f args = else [] in let body, c' = - if defined then Some c', Retyping.get_type_of ctxenv !sigma (EConstr.of_constr c') + if defined then Some c', EConstr.of_constr (Retyping.get_type_of ctxenv !sigma c') else None, c' in let typ = Tacmach.pf_get_hyp_typ gl id in + let typ = EConstr.of_constr typ in let tac = make_abstract_generalize (pf_env gl) id typ concl dep ctx body c' eqs args refls in let tac = Proofview.Unsafe.tclEVARS !sigma <*> tac in Some (tac, dep, succ (List.length ctx), vars) @@ -3743,13 +3790,15 @@ let abstract_generalize ?(generalize_vars=true) ?(force_dep=false) id = let open Context.Named.Declaration in Proofview.Goal.nf_enter { enter = begin fun gl -> Coqlib.check_required_library Coqlib.jmeq_module_name; + let sigma = Tacmach.New.project gl in let (f, args, def, id, oldid) = let oldid = Tacmach.New.pf_get_new_id id gl in match Tacmach.New.pf_get_hyp id gl with - | LocalAssum (_,t) -> let f, args = decompose_app t in + | LocalAssum (_,t) -> let f, args = decompose_app sigma (EConstr.of_constr t) in (f, args, false, id, oldid) | LocalDef (_,t,_) -> - let f, args = decompose_app t in + let t = EConstr.of_constr t in + let f, args = decompose_app sigma t in (f, args, true, id, oldid) in if List.is_empty args then Proofview.tclUNIT () @@ -3778,31 +3827,35 @@ let abstract_generalize ?(generalize_vars=true) ?(force_dep=false) id = Tacticals.New.tclTRY (generalize_dep ~with_let:true (mkVar id))) vars]) end } -let rec compare_upto_variables x y = - if (isVar x || isRel x) && (isVar y || isRel y) then true - else compare_constr compare_upto_variables x y +let compare_upto_variables sigma x y = + let rec compare x y = + if (isVar sigma x || isRel sigma x) && (isVar sigma y || isRel sigma y) then true + else compare_constr sigma compare x y + in + compare x y let specialize_eqs id gl = let open Context.Rel.Declaration in let env = Tacmach.pf_env gl in let ty = Tacmach.pf_get_hyp_typ gl id in + let ty = EConstr.of_constr ty in let evars = ref (project gl) in let unif env evars c1 c2 = - compare_upto_variables c1 c2 && Evarconv.e_conv env evars (EConstr.of_constr c1) (EConstr.of_constr c2) + compare_upto_variables !evars c1 c2 && Evarconv.e_conv env evars c1 c2 in let rec aux in_eqs ctx acc ty = - match kind_of_term ty with + match EConstr.kind !evars ty with | Prod (na, t, b) -> - (match kind_of_term t with - | App (eq, [| eqty; x; y |]) when Term.eq_constr (Lazy.force coq_eq) eq -> - let c = if noccur_between 1 (List.length ctx) x then y else x in + (match EConstr.kind !evars t with + | App (eq, [| eqty; x; y |]) when EConstr.eq_constr !evars (Lazy.force coq_eq) eq -> + let c = if noccur_between !evars 1 (List.length ctx) x then y else x in let pt = mkApp (Lazy.force coq_eq, [| eqty; c; c |]) in let p = mkApp (Lazy.force coq_eq_refl, [| eqty; c |]) in if unif (push_rel_context ctx env) evars pt t then aux true ctx (mkApp (acc, [| p |])) (subst1 p b) else acc, in_eqs, ctx, ty - | App (heq, [| eqty; x; eqty'; y |]) when Term.eq_constr heq (Lazy.force coq_heq) -> - let eqt, c = if noccur_between 1 (List.length ctx) x then eqty', y else eqty, x in + | App (heq, [| eqty; x; eqty'; y |]) when EConstr.eq_constr !evars heq (Lazy.force coq_heq) -> + let eqt, c = if noccur_between !evars 1 (List.length ctx) x then eqty', y else eqty, x in let pt = mkApp (Lazy.force coq_heq, [| eqt; c; eqt; c |]) in let p = mkApp (Lazy.force coq_heq_refl, [| eqt; c |]) in if unif (push_rel_context ctx env) evars pt t then @@ -3811,20 +3864,21 @@ let specialize_eqs id gl = | _ -> if in_eqs then acc, in_eqs, ctx, ty else - let e = e_new_evar (push_rel_context ctx env) evars (EConstr.of_constr t) in - aux false (LocalDef (na,e,t) :: ctx) (mkApp (lift 1 acc, [| mkRel 1 |])) b) + let e = e_new_evar (push_rel_context ctx env) evars t in + aux false (local_def (na,e,t) :: ctx) (mkApp (lift 1 acc, [| mkRel 1 |])) b) | t -> acc, in_eqs, ctx, ty in let acc, worked, ctx, ty = aux false [] (mkVar id) ty in let ctx' = nf_rel_context_evar !evars ctx in let ctx'' = List.map (function - | LocalDef (n,k,t) when isEvar k -> LocalAssum (n,t) + | LocalDef (n,k,t) when isEvar !evars (EConstr.of_constr k) -> LocalAssum (n,t) | decl -> decl) ctx' in let ty' = it_mkProd_or_LetIn ty ctx'' in let acc' = it_mkLambda_or_LetIn acc ctx'' in - let ty' = Tacred.whd_simpl env !evars (EConstr.of_constr ty') - and acc' = Tacred.whd_simpl env !evars (EConstr.of_constr acc') in + let ty' = Tacred.whd_simpl env !evars ty' + and acc' = Tacred.whd_simpl env !evars acc' in + let acc' = EConstr.of_constr acc' in let ty' = Evarutil.nf_evar !evars ty' in let ty' = EConstr.of_constr ty' in if worked then @@ -3840,8 +3894,8 @@ let specialize_eqs id = Proofview.Goal.nf_enter { enter = begin fun gl -> Proofview.V82.tactic (specialize_eqs id) end } -let occur_rel n c = - let res = not (noccurn n c) in +let occur_rel sigma n c = + let res = not (noccurn sigma n c) in res (* This function splits the products of the induction scheme [elimt] into four @@ -3852,20 +3906,20 @@ let occur_rel n c = if there is no branch, we try to fill in acc3 with args/indargs. We also return the conclusion. *) -let decompose_paramspred_branch_args elimt = +let decompose_paramspred_branch_args sigma elimt = let open Context.Rel.Declaration in let rec cut_noccur elimt acc2 = - match kind_of_term elimt with + match EConstr.kind sigma elimt with | Prod(nme,tpe,elimt') -> - let hd_tpe,_ = decompose_app ((strip_prod_assum tpe)) in - if not (occur_rel 1 elimt') && isRel hd_tpe - then cut_noccur elimt' (LocalAssum (nme,tpe)::acc2) - else let acc3,ccl = decompose_prod_assum elimt in acc2 , acc3 , ccl + let hd_tpe,_ = decompose_app sigma (snd (decompose_prod_assum sigma tpe)) in + if not (occur_rel sigma 1 elimt') && isRel sigma hd_tpe + then cut_noccur elimt' (local_assum (nme,tpe)::acc2) + else let acc3,ccl = decompose_prod_assum sigma elimt in acc2 , acc3 , ccl | App(_, _) | Rel _ -> acc2 , [] , elimt | _ -> error_ind_scheme "" in let rec cut_occur elimt acc1 = - match kind_of_term elimt with - | Prod(nme,tpe,c) when occur_rel 1 c -> cut_occur c (LocalAssum (nme,tpe)::acc1) + match EConstr.kind sigma elimt with + | Prod(nme,tpe,c) when occur_rel sigma 1 c -> cut_occur c (local_assum (nme,tpe)::acc1) | Prod(nme,tpe,c) -> let acc2,acc3,ccl = cut_noccur elimt [] in acc1,acc2,acc3,ccl | App(_, _) | Rel _ -> acc1,[],[],elimt | _ -> error_ind_scheme "" in @@ -3878,17 +3932,17 @@ let decompose_paramspred_branch_args elimt = args. We suppose there is only one predicate here. *) match acc2 with | [] -> - let hyps,ccl = decompose_prod_assum elimt in - let hd_ccl_pred,_ = decompose_app ccl in - begin match kind_of_term hd_ccl_pred with + let hyps,ccl = decompose_prod_assum sigma elimt in + let hd_ccl_pred,_ = decompose_app sigma ccl in + begin match EConstr.kind sigma hd_ccl_pred with | Rel i -> let acc3,acc1 = List.chop (i-1) hyps in acc1 , [] , acc3 , ccl | _ -> error_ind_scheme "" end | _ -> acc1, acc2 , acc3, ccl -let exchange_hd_app subst_hd t = - let hd,args= decompose_app t in mkApp (subst_hd,Array.of_list args) +let exchange_hd_app sigma subst_hd t = + let hd,args= decompose_app sigma t in mkApp (subst_hd,Array.of_list args) (* Builds an elim_scheme from its type and calling form (const+binding). We first separate branches. We obtain branches, hyps before (params + preds), @@ -3906,14 +3960,14 @@ let exchange_hd_app subst_hd t = predicates are cited in the conclusion. - finish to fill in the elim_scheme: indarg/farg/args and finally indref. *) -let compute_elim_sig ?elimc elimt = +let compute_elim_sig sigma ?elimc elimt = let open Context.Rel.Declaration in let params_preds,branches,args_indargs,conclusion = - decompose_paramspred_branch_args elimt in + decompose_paramspred_branch_args sigma elimt in - let ccl = exchange_hd_app (mkVar (Id.of_string "__QI_DUMMY__")) conclusion in + let ccl = exchange_hd_app sigma (mkVar (Id.of_string "__QI_DUMMY__")) conclusion in let concl_with_args = it_mkProd_or_LetIn ccl args_indargs in - let nparams = Int.Set.cardinal (free_rels Evd.empty (** FIXME *) (EConstr.of_constr concl_with_args)) in + let nparams = Int.Set.cardinal (free_rels sigma concl_with_args) in let preds,params = List.chop (List.length params_preds - nparams) params_preds in (* A first approximation, further analysis will tweak it *) @@ -3922,7 +3976,7 @@ let compute_elim_sig ?elimc elimt = elimc = elimc; elimt = elimt; concl = conclusion; predicates = preds; npredicates = List.length preds; branches = branches; nbranches = List.length branches; - farg_in_concl = isApp ccl && isApp (last_arg ccl); + farg_in_concl = isApp sigma ccl && isApp sigma (last_arg sigma ccl); params = params; nparams = nparams; (* all other fields are unsure at this point. Including these:*) args = args_indargs; nargs = List.length args_indargs; } in @@ -3943,9 +3997,10 @@ let compute_elim_sig ?elimc elimt = match List.hd args_indargs with | LocalDef (hiname,_,hi) -> error_ind_scheme "" | LocalAssum (hiname,hi) -> - let hi_ind, hi_args = decompose_app hi in + let hi = EConstr.of_constr hi in + let hi_ind, hi_args = decompose_app sigma hi in let hi_is_ind = (* hi est d'un type globalisable *) - match kind_of_term hi_ind with + match EConstr.kind sigma hi_ind with | Ind (mind,_) -> true | Var _ -> true | Const _ -> true @@ -3958,7 +4013,7 @@ let compute_elim_sig ?elimc elimt = else (* Last arg is the indarg *) res := {!res with indarg = Some (List.hd !res.args); - indarg_in_concl = occur_rel 1 ccl; + indarg_in_concl = occur_rel sigma 1 ccl; args = List.tl !res.args; nargs = !res.nargs - 1; }; raise Exit); @@ -3968,55 +4023,58 @@ let compute_elim_sig ?elimc elimt = | None -> !res (* No indref *) | Some (LocalDef _) -> error_ind_scheme "" | Some (LocalAssum (_,ind)) -> - let indhd,indargs = decompose_app ind in - try {!res with indref = Some (global_of_constr indhd) } + let ind = EConstr.of_constr ind in + let indhd,indargs = decompose_app sigma ind in + try {!res with indref = Some (fst (Termops.global_of_constr sigma indhd)) } with e when CErrors.noncritical e -> error "Cannot find the inductive type of the inductive scheme." let compute_scheme_signature evd scheme names_info ind_type_guess = let open Context.Rel.Declaration in - let f,l = decompose_app scheme.concl in + let f,l = decompose_app evd scheme.concl in (* Vérifier que les arguments de Qi sont bien les xi. *) let cond, check_concl = match scheme.indarg with | Some (LocalDef _) -> error "Strange letin, cannot recognize an induction scheme." | None -> (* Non standard scheme *) - let cond hd = Term.eq_constr hd ind_type_guess && not scheme.farg_in_concl + let cond hd = EConstr.eq_constr evd hd ind_type_guess && not scheme.farg_in_concl in (cond, fun _ _ -> ()) | Some (LocalAssum (_,ind)) -> (* Standard scheme from an inductive type *) - let indhd,indargs = decompose_app ind in - let cond hd = Term.eq_constr hd indhd in + let ind = EConstr.of_constr ind in + let indhd,indargs = decompose_app evd ind in + let cond hd = EConstr.eq_constr evd hd indhd in let check_concl is_pred p = (* Check again conclusion *) let ccl_arg_ok = is_pred (p + scheme.nargs + 1) f == IndArg in let ind_is_ok = - List.equal Term.eq_constr + List.equal (fun c1 c2 -> EConstr.eq_constr evd c1 c2) (List.lastn scheme.nargs indargs) - (Context.Rel.to_extended_list 0 scheme.args) in + (List.map EConstr.of_constr (Context.Rel.to_extended_list 0 scheme.args)) in if not (ccl_arg_ok && ind_is_ok) then error_ind_scheme "the conclusion of" in (cond, check_concl) in let is_pred n c = - let hd = fst (decompose_app c) in - match kind_of_term hd with + let hd = fst (decompose_app evd c) in + match EConstr.kind evd hd with | Rel q when n < q && q <= n+scheme.npredicates -> IndArg | _ when cond hd -> RecArg | _ -> OtherArg in let rec check_branch p c = - match kind_of_term c with + match EConstr.kind evd c with | Prod (_,t,c) -> - (is_pred p t, true, not (EConstr.Vars.noccurn evd 1 (EConstr.of_constr c))) :: check_branch (p+1) c + (is_pred p t, true, not (Vars.noccurn evd 1 c)) :: check_branch (p+1) c | LetIn (_,_,_,c) -> - (OtherArg, false, not (EConstr.Vars.noccurn evd 1 (EConstr.of_constr c))) :: check_branch (p+1) c + (OtherArg, false, not (Vars.noccurn evd 1 c)) :: check_branch (p+1) c | _ when is_pred p c == IndArg -> [] | _ -> raise Exit in let rec find_branches p lbrch = match lbrch with | LocalAssum (_,t) :: brs -> + let t = EConstr.of_constr t in (try let lchck_brch = check_branch p t in let n = List.fold_left @@ -4042,7 +4100,7 @@ let compute_scheme_signature evd scheme names_info ind_type_guess = the non standard case, naming of generated hypos is slightly different. *) let compute_elim_signature (evd,(elimc,elimt),ind_type_guess) names_info = - let scheme = compute_elim_sig ~elimc:elimc elimt in + let scheme = compute_elim_sig evd ~elimc:elimc elimt in evd, (compute_scheme_signature evd scheme names_info ind_type_guess, scheme) let guess_elim isrec dep s hyp0 gl = @@ -4057,40 +4115,47 @@ let guess_elim isrec dep s hyp0 gl = if use_dependent_propositions_elimination () && dep then let Sigma (ind, sigma, _) = build_case_analysis_scheme env sigma mind true s in + let ind = EConstr.of_constr ind in (Sigma.to_evar_map sigma, ind) else let Sigma (ind, sigma, _) = build_case_analysis_scheme_default env sigma mind s in + let ind = EConstr.of_constr ind in (Sigma.to_evar_map sigma, ind) in - let elimt = Tacmach.New.pf_unsafe_type_of gl (EConstr.of_constr elimc) in + let elimt = Tacmach.New.pf_unsafe_type_of gl elimc in + let elimt = EConstr.of_constr elimt in evd, ((elimc, NoBindings), elimt), mkIndU mind let given_elim hyp0 (elimc,lbind as e) gl = + let sigma = Tacmach.New.project gl in let tmptyp0 = Tacmach.New.pf_get_hyp_typ hyp0 gl in - let ind_type_guess,_ = decompose_app ((strip_prod tmptyp0)) in - let elimc = EConstr.of_constr elimc in - Tacmach.New.project gl, (e, Tacmach.New.pf_unsafe_type_of gl elimc), ind_type_guess + let tmptyp0 = EConstr.of_constr tmptyp0 in + let ind_type_guess,_ = decompose_app sigma (snd (decompose_prod sigma tmptyp0)) in + let elimt = Tacmach.New.pf_unsafe_type_of gl elimc in + let elimt = EConstr.of_constr elimt in + Tacmach.New.project gl, (e, elimt), ind_type_guess type scheme_signature = (Id.t list * (elim_arg_kind * bool * bool * Id.t) list) array type eliminator_source = - | ElimUsing of (eliminator * types) * scheme_signature + | ElimUsing of (eliminator * EConstr.types) * scheme_signature | ElimOver of bool * Id.t let find_induction_type isrec elim hyp0 gl = + let sigma = Tacmach.New.project gl in let scheme,elim = match elim with | None -> let sort = Tacticals.New.elimination_sort_of_goal gl in let _, (elimc,elimt),_ = guess_elim isrec (* dummy: *) true sort hyp0 gl in - let scheme = compute_elim_sig ~elimc elimt in + let scheme = compute_elim_sig sigma ~elimc elimt in (* We drop the scheme waiting to know if it is dependent *) scheme, ElimOver (isrec,hyp0) | Some e -> let evd, (elimc,elimt),ind_guess = given_elim hyp0 e gl in - let scheme = compute_elim_sig ~elimc elimt in + let scheme = compute_elim_sig sigma ~elimc elimt in if Option.is_empty scheme.indarg then error "Cannot find induction type"; let indsign = compute_scheme_signature evd scheme hyp0 ind_guess in let elim = ({elimindex = Some(-1); elimbody = elimc; elimrename = None},elimt) in @@ -4104,7 +4169,8 @@ let get_elim_signature elim hyp0 gl = compute_elim_signature (given_elim hyp0 elim gl) hyp0 let is_functional_induction elimc gl = - let scheme = compute_elim_sig ~elimc (Tacmach.New.pf_unsafe_type_of gl (EConstr.of_constr (fst elimc))) in + let sigma = Tacmach.New.project gl in + let scheme = compute_elim_sig sigma ~elimc (EConstr.of_constr (Tacmach.New.pf_unsafe_type_of gl (fst elimc))) in (* The test is not safe: with non-functional induction on non-standard induction scheme, this may fail *) Option.is_empty scheme.indarg @@ -4128,17 +4194,18 @@ let get_eliminator elim dep s gl = of lid are parameters (first ones), the other are arguments. Returns the clause obtained. *) let recolle_clenv i params args elimclause gl = - let _,arr = destApp (EConstr.Unsafe.to_constr elimclause.templval.rebus) in + let _,arr = destApp elimclause.evd elimclause.templval.rebus in let lindmv = Array.map (fun x -> - match kind_of_term x with + match EConstr.kind elimclause.evd x with | Meta mv -> mv | _ -> user_err ~hdr:"elimination_clause" (str "The type of the elimination clause is not well-formed.")) arr in let k = match i with -1 -> Array.length lindmv - List.length args | _ -> i in (* parameters correspond to first elts of lid. *) + let pf_get_hyp_typ id gl = EConstr.of_constr (pf_get_hyp_typ id gl) in let clauses_params = List.map_i (fun i id -> mkVar id , pf_get_hyp_typ id gl, lindmv.(i)) 0 params in @@ -4153,8 +4220,6 @@ let recolle_clenv i params args elimclause gl = (* from_n (Some 0) means that x should be taken "as is" without trying to unify (which would lead to trying to apply it to evars if y is a product). *) - let x = EConstr.of_constr x in - let y = EConstr.of_constr y in let indclause = Tacmach.New.of_old (fun gl -> mk_clenv_from_n gl (Some 0) (x,y)) gl in let elimclause' = clenv_fchain ~with_univs:false i acc indclause in elimclause') @@ -4167,14 +4232,12 @@ let recolle_clenv i params args elimclause gl = *) let induction_tac with_evars params indvars elim = Proofview.Goal.nf_enter { enter = begin fun gl -> + let sigma = Tacmach.New.project gl in let ({elimindex=i;elimbody=(elimc,lbindelimc);elimrename=rename},elimt) = elim in - let i = match i with None -> index_of_ind_arg elimt | Some i -> i in + let i = match i with None -> index_of_ind_arg sigma elimt | Some i -> i in (* elimclause contains this: (elimc ?i ?j ?k...?l) *) - let elimc = contract_letin_in_lam_header elimc in + let elimc = contract_letin_in_lam_header sigma elimc in let elimc = mkCast (elimc, DEFAULTcast, elimt) in - let elimc = EConstr.of_constr elimc in - let elimt = EConstr.of_constr elimt in - let lbindelimc = Miscops.map_bindings EConstr.of_constr lbindelimc in let elimclause = pf_apply make_clenv_binding gl (elimc,elimt) lbindelimc in (* elimclause' is built from elimclause by instanciating all args and params. *) let elimclause' = recolle_clenv i params indvars elimclause gl in @@ -4197,7 +4260,8 @@ let apply_induction_in_context with_evars hyp0 inhyps elim indvars names induct_ let dep_in_concl = Option.cata (fun id -> occur_var env sigma id (EConstr.of_constr concl)) false hyp0 in let dep = dep_in_hyps || dep_in_concl in let tmpcl = it_mkNamedProd_or_LetIn concl deps in - let s = Retyping.get_sort_family_of env sigma (EConstr.of_constr tmpcl) in + let tmpcl = EConstr.of_constr tmpcl in + let s = Retyping.get_sort_family_of env sigma tmpcl in let deps_cstr = List.fold_left (fun a decl -> if NamedDecl.is_local_assum decl then (mkVar (NamedDecl.get_id decl))::a else a) [] deps in @@ -4321,14 +4385,12 @@ let use_bindings env sigma elim must_be_closed (c,lbind) typ = let rec find_clause typ = try let typ = EConstr.of_constr typ in - let c = EConstr.of_constr c in - let lbind = Miscops.map_bindings EConstr.of_constr lbind in let indclause = make_clenv_binding env sigma (c,typ) lbind in if must_be_closed && occur_meta indclause.evd (clenv_value indclause) then error "Need a fully applied argument."; (* We lose the possibility of coercions in with-bindings *) let (sigma, c) = pose_all_metas_as_evars env indclause.evd (clenv_value indclause) in - Sigma.Unsafe.of_pair (EConstr.Unsafe.to_constr c, sigma) + Sigma.Unsafe.of_pair (c, sigma) with e when catchable_exception e -> try find_clause (try_red_product env sigma (EConstr.of_constr typ)) with Redelimination -> raise e in @@ -4337,8 +4399,6 @@ let use_bindings env sigma elim must_be_closed (c,lbind) typ = let check_expected_type env sigma (elimc,bl) elimt = (* Compute the expected template type of the term in case a using clause is given *) - let open EConstr in - let elimt = EConstr.of_constr elimt in let sign,_ = splay_prod env sigma elimt in let n = List.length sign in if n == 0 then error "Scheme cannot be applied."; @@ -4354,11 +4414,11 @@ let check_enough_applied env sigma elim = | None -> (* No eliminator given *) fun u -> - let t,_ = decompose_app (whd_all env sigma u) in isInd t + let t,_ = decompose_app sigma (EConstr.of_constr (whd_all env sigma u)) in isInd sigma t | Some elimc -> - let elimt = Retyping.get_type_of env sigma (EConstr.of_constr (fst elimc)) in - let scheme = compute_elim_sig ~elimc elimt in - let elimc = Miscops.map_with_bindings EConstr.of_constr elimc in + let elimt = Retyping.get_type_of env sigma (fst elimc) in + let elimt = EConstr.of_constr elimt in + let scheme = compute_elim_sig sigma ~elimc elimt in match scheme.indref with | None -> (* in the absence of information, do not assume it may be @@ -4381,11 +4441,9 @@ let pose_induction_arg_then isrec with_evars (is_arg_pure_hyp,from_prefix) elim let store = Proofview.Goal.extra gl in let check = check_enough_applied env sigma elim in let Sigma (c, sigma', p) = use_bindings env sigma elim false (c0,lbind) t0 in - let c = EConstr.of_constr c in let abs = AbstractPattern (from_prefix,check,Name id,(pending,c),cls,false) in let ccl = EConstr.of_constr ccl in let (id,sign,_,lastlhyp,ccl,res) = make_abstraction env sigma' ccl abs in - let ccl = EConstr.Unsafe.to_constr ccl in match res with | None -> (* pattern not found *) @@ -4393,9 +4451,7 @@ let pose_induction_arg_then isrec with_evars (is_arg_pure_hyp,from_prefix) elim (* we restart using bindings after having tried type-class resolution etc. on the term given by the user *) let flags = tactic_infer_flags (with_evars && (* do not give a success semantics to edestruct on an open term yet *) false) in - let c0 = EConstr.of_constr c0 in let Sigma (c0, sigma, q) = finish_evar_resolution ~flags env sigma (pending,c0) in - let c0 = EConstr.Unsafe.to_constr c0 in let tac = (if isrec then (* Historically, induction has side conditions last *) @@ -4407,13 +4463,14 @@ let pose_induction_arg_then isrec with_evars (is_arg_pure_hyp,from_prefix) elim Refine.refine ~unsafe:true { run = begin fun sigma -> let b = not with_evars && with_eq != None in let Sigma (c, sigma, p) = use_bindings env sigma elim b (c0,lbind) t0 in - let t = Retyping.get_type_of env (Sigma.to_evar_map sigma) (EConstr.of_constr c) in + let t = Retyping.get_type_of env (Sigma.to_evar_map sigma) c in + let t = EConstr.of_constr t in let Sigma (ans, sigma, q) = mkletin_goal env sigma store with_eq false (id,lastlhyp,ccl,c) (Some t) in Sigma (ans, sigma, p +> q) end }; if with_evars then Proofview.shelve_unifiable else guard_no_unifiable; if is_arg_pure_hyp - then Tacticals.New.tclTRY (clear [destVar c0]) + then Proofview.tclEVARMAP >>= fun sigma -> Tacticals.New.tclTRY (clear [destVar sigma c0]) else Proofview.tclUNIT (); if isrec then Proofview.cycle (-1) else Proofview.tclUNIT () ]) @@ -4422,7 +4479,6 @@ let pose_induction_arg_then isrec with_evars (is_arg_pure_hyp,from_prefix) elim Sigma (tac, sigma, q) | Some (Sigma (c, sigma', q)) -> - let c = EConstr.Unsafe.to_constr c in (* pattern found *) let with_eq = Option.map (fun eq -> (false,eq)) eqname in (* TODO: if ind has predicate parameters, use JMeq instead of eq *) @@ -4451,14 +4507,15 @@ let induction_gen clear_flag isrec with_evars elim Proofview.Goal.enter { enter = begin fun gl -> let env = Proofview.Goal.env gl in let sigma = Proofview.Goal.sigma gl in + let evd = Sigma.to_evar_map sigma in let ccl = Proofview.Goal.raw_concl gl in let cls = Option.default allHypsAndConcl cls in - let t = typ_of env sigma (EConstr.of_constr c) in + let t = typ_of env sigma c in let is_arg_pure_hyp = - isVar c && not (mem_named_context_val (destVar c) (Global.named_context_val ())) + isVar evd c && not (mem_named_context_val (destVar evd c) (Global.named_context_val ())) && lbind == NoBindings && not with_evars && Option.is_empty eqname && clear_flag == None - && has_generic_occurrences_but_goal cls (destVar c) env (Sigma.to_evar_map sigma) ccl in + && has_generic_occurrences_but_goal cls (destVar evd c) env evd ccl in let enough_applied = check_enough_applied env sigma elim (EConstr.of_constr t) in if is_arg_pure_hyp && enough_applied then (* First case: induction on a variable already in an inductive type and @@ -4466,7 +4523,7 @@ let induction_gen clear_flag isrec with_evars elim This is a situation where the induction argument is a clearable variable of the goal w/o occurrence selection and w/o equality kept: no need to generalize *) - let id = destVar c in + let id = destVar evd c in Tacticals.New.tclTHEN (clear_unselected_context id inhyps cls) (induction_with_atomization_of_ind_arg @@ -4501,7 +4558,8 @@ let induction_gen_l isrec with_evars elim names lc = match l with | [] -> Proofview.tclUNIT () | c::l' -> - match kind_of_term c with + Proofview.tclEVARMAP >>= fun sigma -> + match EConstr.kind sigma c with | Var id when not (mem_named_context_val id (Global.named_context_val ())) && not with_evars -> let _ = newlc:= id::!newlc in @@ -4512,10 +4570,10 @@ let induction_gen_l isrec with_evars elim names lc = let type_of = Tacmach.New.pf_unsafe_type_of gl in let sigma = Tacmach.New.project gl in let x = - id_of_name_using_hdchar (Global.env()) (type_of (EConstr.of_constr c)) Anonymous in + id_of_name_using_hdchar (Global.env()) (type_of c) Anonymous in let id = new_fresh_id [] x gl in - let newl' = List.map (fun r -> replace_term sigma (EConstr.of_constr c) (EConstr.mkVar id) (EConstr.of_constr r)) l' in + let newl' = List.map (fun r -> EConstr.of_constr (replace_term sigma c (mkVar id) r)) l' in let _ = newlc:=id::!newlc in Tacticals.New.tclTHEN (letin_tac None (Name id) c None allHypsAndConcl) @@ -4639,13 +4697,12 @@ let simple_destruct = function let elim_scheme_type elim t = Proofview.Goal.nf_enter { enter = begin fun gl -> - let elim = EConstr.of_constr elim in let clause = Tacmach.New.of_old (fun gl -> mk_clenv_type_of gl elim) gl in - match kind_of_term (last_arg (EConstr.Unsafe.to_constr clause.templval.rebus)) with + match EConstr.kind clause.evd (last_arg clause.evd clause.templval.rebus) with | Meta mv -> let clause' = (* t is inductive, then CUMUL or CONV is irrelevant *) - clenv_unify ~flags:(elim_flags ()) Reduction.CUMUL (EConstr.of_constr t) + clenv_unify ~flags:(elim_flags ()) Reduction.CUMUL t (clenv_meta_type clause mv) clause in Clenvtac.res_pf clause' ~flags:(elim_flags ()) ~with_evars:false | _ -> anomaly (Pp.str "elim_scheme_type") @@ -4653,7 +4710,6 @@ let elim_scheme_type elim t = let elim_type t = Proofview.Goal.s_enter { s_enter = begin fun gl -> - let t = EConstr.of_constr t in let (ind,t) = Tacmach.New.pf_apply reduce_to_atomic_ind gl t in let evd, elimc = find_ind_eliminator (fst ind) (Tacticals.New.elimination_sort_of_goal gl) gl in Sigma.Unsafe.of_pair (elim_scheme_type elimc t, evd) @@ -4661,12 +4717,12 @@ let elim_type t = let case_type t = Proofview.Goal.s_enter { s_enter = begin fun gl -> - let t = EConstr.of_constr t in let sigma = Proofview.Goal.sigma gl in let env = Tacmach.New.pf_env gl in let (ind,t) = reduce_to_atomic_ind env (Sigma.to_evar_map sigma) t in let s = Tacticals.New.elimination_sort_of_goal gl in let Sigma (elimc, evd, p) = build_case_analysis_scheme_default env sigma ind s in + let elimc = EConstr.of_constr elimc in Sigma (elim_scheme_type elimc t, evd, p) end } @@ -4722,12 +4778,10 @@ let (forward_setoid_symmetry, setoid_symmetry) = Hook.make () (* This is probably not very useful any longer *) let prove_symmetry hdcncl eq_kind = let symc = - let open EConstr in match eq_kind with | MonomorphicLeibnizEq (c1,c2) -> mkApp(hdcncl,[|c2;c1|]) | PolymorphicLeibnizEq (typ,c1,c2) -> mkApp(hdcncl,[|typ;c2;c1|]) | HeterogenousEq (t1,c1,t2,c2) -> mkApp(hdcncl,[|t2;c2;t1;c1|]) in - let symc = EConstr.Unsafe.to_constr symc in Tacticals.New.tclTHENFIRST (cut symc) (Tacticals.New.tclTHENLIST [ intro; @@ -4748,12 +4802,13 @@ let symmetry_red allowred = inside setoid_reflexivity (see Optimize cases in setoid_replace.ml). *) let sigma = Tacmach.New.project gl in let concl = maybe_betadeltaiota_concl allowred gl in - match_with_equation sigma (EConstr.of_constr concl) >>= fun with_eqn -> + let concl = EConstr.of_constr concl in + match_with_equation sigma concl >>= fun with_eqn -> match with_eqn with | Some eq_data,_,_ -> Tacticals.New.tclTHEN (convert_concl_no_check concl DEFAULTcast) - (Tacticals.New.pf_constr_of_global eq_data.sym apply) + (Tacticals.New.pf_constr_of_global eq_data.sym (EConstr.of_constr %> apply)) | None,eq,eq_kind -> prove_symmetry eq eq_kind end } @@ -4771,20 +4826,18 @@ let (forward_setoid_symmetry_in, setoid_symmetry_in) = Hook.make () let symmetry_in id = Proofview.Goal.enter { enter = begin fun gl -> let sigma = Tacmach.New.project gl in - let ctype = Tacmach.New.pf_unsafe_type_of gl (EConstr.mkVar id) in - let sign,t = decompose_prod_assum ctype in - let t = EConstr.of_constr t in + let ctype = Tacmach.New.pf_unsafe_type_of gl (mkVar id) in + let ctype = EConstr.of_constr ctype in + let sign,t = decompose_prod_assum sigma ctype in Proofview.tclORELSE begin match_with_equation sigma t >>= fun (_,hdcncl,eq) -> let symccl = - let open EConstr in match eq with | MonomorphicLeibnizEq (c1,c2) -> mkApp (hdcncl, [| c2; c1 |]) | PolymorphicLeibnizEq (typ,c1,c2) -> mkApp (hdcncl, [| typ; c2; c1 |]) | HeterogenousEq (t1,c1,t2,c2) -> mkApp (hdcncl, [| t2; c2; t1; c1 |]) in - let symccl = EConstr.Unsafe.to_constr symccl in - Tacticals.New.tclTHENS (cut (it_mkProd_or_LetIn symccl sign)) + Tacticals.New.tclTHENS (cut (EConstr.it_mkProd_or_LetIn symccl sign)) [ intro_replacing id; Tacticals.New.tclTHENLIST [ intros; symmetry; apply (mkVar id); assumption ] ] end @@ -4818,8 +4871,6 @@ let (forward_setoid_transitivity, setoid_transitivity) = Hook.make () (* This is probably not very useful any longer *) let prove_transitivity hdcncl eq_kind t = Proofview.Goal.enter { enter = begin fun gl -> - let t = EConstr.of_constr t in - let open EConstr in let (eq1,eq2) = match eq_kind with | MonomorphicLeibnizEq (c1,c2) -> mkApp (hdcncl, [| c1; t|]), mkApp (hdcncl, [| t; c2 |]) @@ -4834,8 +4885,6 @@ let prove_transitivity hdcncl eq_kind t = (mkApp(hdcncl, [| typ1; c1; typt ;t |]), mkApp(hdcncl, [| typt; t; typ2; c2 |])) in - let eq1 = EConstr.Unsafe.to_constr eq1 in - let eq2 = EConstr.Unsafe.to_constr eq2 in Tacticals.New.tclTHENFIRST (cut eq2) (Tacticals.New.tclTHENFIRST (cut eq1) (Tacticals.New.tclTHENLIST @@ -4851,14 +4900,15 @@ let transitivity_red allowred t = inside setoid_reflexivity (see Optimize cases in setoid_replace.ml). *) let sigma = Tacmach.New.project gl in let concl = maybe_betadeltaiota_concl allowred gl in - match_with_equation sigma (EConstr.of_constr concl) >>= fun with_eqn -> + let concl = EConstr.of_constr concl in + match_with_equation sigma concl >>= fun with_eqn -> match with_eqn with | Some eq_data,_,_ -> Tacticals.New.tclTHEN (convert_concl_no_check concl DEFAULTcast) (match t with - | None -> Tacticals.New.pf_constr_of_global eq_data.trans eapply - | Some t -> Tacticals.New.pf_constr_of_global eq_data.trans (fun trans -> apply_list [trans;t])) + | None -> Tacticals.New.pf_constr_of_global eq_data.trans (EConstr.of_constr %> eapply) + | Some t -> Tacticals.New.pf_constr_of_global eq_data.trans (fun trans -> apply_list [EConstr.of_constr trans;t])) | None,eq,eq_kind -> match t with | None -> Tacticals.New.tclZEROMSG (str"etransitivity not supported for this relation.") @@ -4902,6 +4952,8 @@ let rec decompose len c t accu = | _ -> assert false let rec shrink ctx sign c t accu = + let open Term in + let open CVars in match ctx, sign with | [], [] -> (c, t, accu) | p :: ctx, decl :: sign -> @@ -4984,6 +5036,7 @@ let abstract_subproof id gk tac = if !shrink_abstract then shrink_entry sign const else (const, List.rev (Context.Named.to_instance sign)) in + let args = List.map EConstr.of_constr args in let cd = Entries.DefinitionEntry const in let decl = (cd, IsProof Lemma) in let cst () = @@ -4995,6 +5048,7 @@ let abstract_subproof id gk tac = let cst = Impargs.with_implicit_protection cst () in (* let evd, lem = Evd.fresh_global (Global.env ()) evd (ConstRef cst) in *) let lem, ctx = Universes.unsafe_constr_of_global (ConstRef cst) in + let lem = EConstr.of_constr lem in let evd = Evd.set_universe_context evd ectx in let open Safe_typing in let eff = private_con_of_con (Global.safe_env ()) cst in @@ -5026,8 +5080,6 @@ let tclABSTRACT name_op tac = abstract_subproof s gk tac let unify ?(state=full_transparent_state) x y = - let x = EConstr.of_constr x in - let y = EConstr.of_constr y in Proofview.Goal.nf_s_enter { s_enter = begin fun gl -> let sigma = Proofview.Goal.sigma gl in try @@ -5081,10 +5133,6 @@ module New = struct {onhyps=None; concl_occs=AllOccurrences } let refine ?unsafe c = - let c = { run = begin fun sigma -> - let Sigma (c, sigma, p) = c.run sigma in - Sigma (EConstr.of_constr c, sigma, p) - end } in Refine.refine ?unsafe c <*> reduce_after_refine end diff --git a/tactics/tactics.mli b/tactics/tactics.mli index 368a1df76..630c660a1 100644 --- a/tactics/tactics.mli +++ b/tactics/tactics.mli @@ -9,6 +9,7 @@ open Loc open Names open Term +open EConstr open Environ open Proof_type open Evd @@ -128,7 +129,7 @@ val exact_proof : Constrexpr.constr_expr -> unit Proofview.tactic (** {6 Reduction tactics. } *) -type tactic_reduction = env -> evar_map -> EConstr.t -> constr +type tactic_reduction = env -> evar_map -> constr -> Constr.constr type change_arg = patvar_map -> constr Sigma.run @@ -259,7 +260,7 @@ type elim_scheme = { farg_in_concl: bool; (** true if (f...) appears at the end of conclusion *) } -val compute_elim_sig : ?elimc: constr with_bindings -> types -> elim_scheme +val compute_elim_sig : evar_map -> ?elimc:constr with_bindings -> types -> elim_scheme (** elim principle with the index of its inductive arg *) type eliminator = { @@ -413,7 +414,7 @@ val subst_one : val declare_intro_decomp_eq : ((int -> unit Proofview.tactic) -> Coqlib.coq_eq_data * types * - (EConstr.types * EConstr.constr * EConstr.constr) -> + (types * constr * constr) -> constr * types -> unit Proofview.tactic) -> unit (** {6 Simple form of basic tactics. } *) diff --git a/tactics/term_dnet.ml b/tactics/term_dnet.ml index 6294f9fdc..38342b64d 100644 --- a/tactics/term_dnet.ml +++ b/tactics/term_dnet.ml @@ -355,7 +355,7 @@ struct with Invalid_argument _ -> [],c_id in let wc,whole_c = if Opt.direction then whole_c,wc else wc,whole_c in try - let _ = Termops.filtering ctx Reduction.CUMUL wc whole_c in + let _ = Termops.filtering Evd.empty ctx Reduction.CUMUL wc whole_c in id :: acc with Termops.CannotFilter -> (* msgnl(str"recon "++Termops.print_constr_env (Global.env()) wc); *) acc ) (TDnet.find_match dpat dn) [] diff --git a/toplevel/auto_ind_decl.ml b/toplevel/auto_ind_decl.ml index 6561627f6..7759c400c 100644 --- a/toplevel/auto_ind_decl.ml +++ b/toplevel/auto_ind_decl.ml @@ -386,9 +386,10 @@ let do_replace_lb mode lb_scheme_key aavoid narg p q = (Array.map (fun x -> x) v) (Array.map (fun x -> do_arg x 1) v)) (Array.map (fun x -> do_arg x 2) v) - in let app = if Array.equal eq_constr lb_args [||] + in let app = if Array.equal Term.eq_constr lb_args [||] then lb_type_of_p else mkApp (lb_type_of_p,lb_args) in + let app = EConstr.of_constr app in Tacticals.New.tclTHENLIST [ Proofview.tclEFFECTS eff; Equality.replace p q ; apply app ; Auto.default_auto] @@ -426,7 +427,7 @@ let do_replace_bl mode bl_scheme_key (ind,u as indu) aavoid narg lft rgt = | (t1::q1,t2::q2) -> Proofview.Goal.enter { enter = begin fun gl -> let tt1 = Tacmach.New.pf_unsafe_type_of gl (EConstr.of_constr t1) in - if eq_constr t1 t2 then aux q1 q2 + if Term.eq_constr t1 t2 then aux q1 q2 else ( let u,v = try destruct_ind tt1 (* trick so that the good sequence is returned*) @@ -455,9 +456,10 @@ let do_replace_bl mode bl_scheme_key (ind,u as indu) aavoid narg lft rgt = (Array.map (fun x -> do_arg x 1) v)) (Array.map (fun x -> do_arg x 2) v ) in - let app = if Array.equal eq_constr bl_args [||] + let app = if Array.equal Term.eq_constr bl_args [||] then bl_t1 else mkApp (bl_t1,bl_args) in + let app = EConstr.of_constr app in Tacticals.New.tclTHENLIST [ Proofview.tclEFFECTS eff; Equality.replace_by t1 t2 @@ -515,7 +517,7 @@ let eqI ind l = try let c, eff = find_scheme beq_scheme_kind ind in mkConst c, eff with Not_found -> user_err ~hdr:"AutoIndDecl.eqI" (str "The boolean equality on " ++ pr_mind (fst ind) ++ str " is needed."); - in (if Array.equal eq_constr eA [||] then e else mkApp(e,eA)), eff + in (if Array.equal Term.eq_constr eA [||] then e else mkApp(e,eA)), eff (**********************************************************************) (* Boolean->Leibniz *) @@ -580,9 +582,9 @@ let compute_bl_tact mode bl_scheme_key ind lnamesparrec nparrec = (* try with *) Tacticals.New.tclTHENLIST [ intros_using fresh_first_intros; intro_using freshn ; - induct_on (mkVar freshn); + induct_on (EConstr.mkVar freshn); intro_using freshm; - destruct_on (mkVar freshm); + destruct_on (EConstr.mkVar freshm); intro_using freshz; intros; Tacticals.New.tclTRY ( @@ -594,10 +596,10 @@ repeat ( apply andb_prop in z;let z1:= fresh "Z" in destruct z as [z1 z]). *) Tacticals.New.tclREPEAT ( Tacticals.New.tclTHENLIST [ - Simple.apply_in freshz (andb_prop()); + Simple.apply_in freshz (EConstr.of_constr (andb_prop())); Proofview.Goal.nf_enter { enter = begin fun gl -> let fresht = fresh_id (Id.of_string "Z") gl in - destruct_on_as (mkVar freshz) + destruct_on_as (EConstr.mkVar freshz) (IntroOrPattern [[dl,IntroNaming (IntroIdentifier fresht); dl,IntroNaming (IntroIdentifier freshz)]]) end } @@ -723,19 +725,19 @@ let compute_lb_tact mode lb_scheme_key ind lnamesparrec nparrec = (* try with *) Tacticals.New.tclTHENLIST [ intros_using fresh_first_intros; intro_using freshn ; - induct_on (mkVar freshn); + induct_on (EConstr.mkVar freshn); intro_using freshm; - destruct_on (mkVar freshm); + destruct_on (EConstr.mkVar freshm); intro_using freshz; intros; Tacticals.New.tclTRY ( Tacticals.New.tclORELSE reflexivity (Equality.discr_tac false None) ); - Equality.inj None false None (mkVar freshz,NoBindings); + Equality.inj None false None (EConstr.mkVar freshz,NoBindings); intros; simpl_in_concl; Auto.default_auto; Tacticals.New.tclREPEAT ( - Tacticals.New.tclTHENLIST [apply (andb_true_intro()); + Tacticals.New.tclTHENLIST [apply (EConstr.of_constr (andb_true_intro())); simplest_split ;Auto.default_auto ] ); Proofview.Goal.nf_enter { enter = begin fun gls -> @@ -888,18 +890,18 @@ let compute_dec_tact ind lnamesparrec nparrec = intros_using fresh_first_intros; intros_using [freshn;freshm]; (*we do this so we don't have to prove the same goal twice *) - assert_by (Name freshH) ( + assert_by (Name freshH) (EConstr.of_constr ( mkApp(sumbool(),[|eqtrue eqbnm; eqfalse eqbnm|]) - ) - (Tacticals.New.tclTHEN (destruct_on eqbnm) Auto.default_auto); + )) + (Tacticals.New.tclTHEN (destruct_on (EConstr.of_constr eqbnm)) Auto.default_auto); Proofview.Goal.nf_enter { enter = begin fun gl -> let freshH2 = fresh_id (Id.of_string "H") gl in - Tacticals.New.tclTHENS (destruct_on_using (mkVar freshH) freshH2) [ + Tacticals.New.tclTHENS (destruct_on_using (EConstr.mkVar freshH) freshH2) [ (* left *) Tacticals.New.tclTHENLIST [ simplest_left; - apply (mkApp(blI,Array.map(fun x->mkVar x) xargs)); + apply (EConstr.of_constr (mkApp(blI,Array.map(fun x->mkVar x) xargs))); Auto.default_auto ] ; @@ -913,9 +915,9 @@ let compute_dec_tact ind lnamesparrec nparrec = intro; Equality.subst_all (); assert_by (Name freshH3) - (mkApp(eq,[|bb;mkApp(eqI,[|mkVar freshm;mkVar freshm|]);tt|])) + (EConstr.of_constr (mkApp(eq,[|bb;mkApp(eqI,[|mkVar freshm;mkVar freshm|]);tt|]))) (Tacticals.New.tclTHENLIST [ - apply (mkApp(lbI,Array.map (fun x->mkVar x) xargs)); + apply (EConstr.of_constr (mkApp(lbI,Array.map (fun x->mkVar x) xargs))); Auto.default_auto ]); Equality.general_rewrite_bindings_in true diff --git a/toplevel/command.ml b/toplevel/command.ml index 80f3b26e4..08f3ad4a7 100644 --- a/toplevel/command.ml +++ b/toplevel/command.ml @@ -1030,8 +1030,8 @@ let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation = let def = mkApp (Universes.constr_of_global (delayed_force fix_sub_ref), [| argtyp ; wf_rel ; - Evarutil.e_new_evar env evdref - ~src:(Loc.ghost, Evar_kinds.QuestionMark (Evar_kinds.Define false)) (EConstr.of_constr wf_proof); + EConstr.Unsafe.to_constr (Evarutil.e_new_evar env evdref + ~src:(Loc.ghost, Evar_kinds.QuestionMark (Evar_kinds.Define false)) (EConstr.of_constr wf_proof)); prop |]) in let def = Typing.e_solve_evars env evdref (EConstr.of_constr def) in -- cgit v1.2.3 From 8b660087beb2209e52bc4412dc82c6727963c6a5 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 18 Nov 2016 20:26:15 +0100 Subject: Elim API using EConstr. --- ltac/extratactics.ml4 | 6 +++--- plugins/funind/functional_principles_proofs.ml | 2 +- plugins/funind/recdef.ml | 2 +- tactics/elim.ml | 19 ++++++++++--------- tactics/elim.mli | 1 + tactics/tacticals.ml | 3 ++- tactics/tacticals.mli | 4 ++-- 7 files changed, 20 insertions(+), 17 deletions(-) diff --git a/ltac/extratactics.ml4 b/ltac/extratactics.ml4 index c39b1a0e9..faf545d4f 100644 --- a/ltac/extratactics.ml4 +++ b/ltac/extratactics.ml4 @@ -172,11 +172,11 @@ END (* Decompose *) TACTIC EXTEND decompose_sum -| [ "decompose" "sum" constr(c) ] -> [ Elim.h_decompose_or c ] +| [ "decompose" "sum" constr(c) ] -> [ Elim.h_decompose_or (EConstr.of_constr c) ] END TACTIC EXTEND decompose_record -| [ "decompose" "record" constr(c) ] -> [ Elim.h_decompose_and c ] +| [ "decompose" "record" constr(c) ] -> [ Elim.h_decompose_and (EConstr.of_constr c) ] END (**********************************************************************) @@ -1055,7 +1055,7 @@ let decompose l c = end } TACTIC EXTEND decompose -| [ "decompose" "[" ne_constr_list(l) "]" constr(c) ] -> [ decompose l c ] +| [ "decompose" "[" ne_constr_list(l) "]" constr(c) ] -> [ decompose l (EConstr.of_constr c) ] END (** library/keys *) diff --git a/plugins/funind/functional_principles_proofs.ml b/plugins/funind/functional_principles_proofs.ml index 503cafdd5..940f1669a 100644 --- a/plugins/funind/functional_principles_proofs.ml +++ b/plugins/funind/functional_principles_proofs.ml @@ -1626,7 +1626,7 @@ let prove_principle_for_gen [ Proofview.V82.of_tactic (generalize [lemma]); Proofview.V82.of_tactic (Simple.intro hid); - Proofview.V82.of_tactic (Elim.h_decompose_and (mkVar hid)); + Proofview.V82.of_tactic (Elim.h_decompose_and (EConstr.mkVar hid)); (fun g -> let new_hyps = pf_ids_of_hyps g in tcc_list := List.rev (List.subtract Id.equal new_hyps (hid::hyps)); diff --git a/plugins/funind/recdef.ml b/plugins/funind/recdef.ml index d5ee42af8..74affa396 100644 --- a/plugins/funind/recdef.ml +++ b/plugins/funind/recdef.ml @@ -1314,7 +1314,7 @@ let open_new_goal build_proof sigma using_lemmas ref_ goal_name (gls_type,decomp (fun g -> let ids = pf_ids_of_hyps g in tclTHEN - (Proofview.V82.of_tactic (Elim.h_decompose_and (mkVar hid))) + (Proofview.V82.of_tactic (Elim.h_decompose_and (EConstr.mkVar hid))) (fun g -> let ids' = pf_ids_of_hyps g in lid := List.rev (List.subtract Id.equal ids' ids); diff --git a/tactics/elim.ml b/tactics/elim.ml index e641f970a..ef848c2e1 100644 --- a/tactics/elim.ml +++ b/tactics/elim.ml @@ -10,6 +10,7 @@ open Util open Names open Term open Termops +open EConstr open Inductiveops open Hipattern open Tacmach.New @@ -55,7 +56,7 @@ Another example : *) let elimHypThen tac id = - elimination_then tac (EConstr.mkVar id) + elimination_then tac (mkVar id) let rec general_decompose_on_hyp recognizer = ifOnHyp recognizer (general_decompose_aux recognizer) (fun _ -> Proofview.tclUNIT()) @@ -77,7 +78,6 @@ let tmphyp_name = Id.of_string "_TmpHyp" let up_to_delta = ref false (* true *) let general_decompose recognizer c = - let c = EConstr.of_constr c in Proofview.Goal.enter { enter = begin fun gl -> let type_of = pf_unsafe_type_of gl in let sigma = project gl in @@ -105,17 +105,17 @@ let head_in indl t gl = let decompose_these c l = Proofview.Goal.enter { enter = begin fun gl -> let indl = List.map (fun x -> x, Univ.Instance.empty) l in - general_decompose (fun sigma (_,t) -> head_in indl (EConstr.of_constr t) gl) c + general_decompose (fun sigma (_,t) -> head_in indl t gl) c end } let decompose_and c = general_decompose - (fun sigma (_,t) -> is_record sigma (EConstr.of_constr t)) + (fun sigma (_,t) -> is_record sigma t) c let decompose_or c = general_decompose - (fun sigma (_,t) -> is_disjunction sigma (EConstr.of_constr t)) + (fun sigma (_,t) -> is_disjunction sigma t) c let h_decompose l c = decompose_these c l @@ -127,7 +127,7 @@ let h_decompose_and = decompose_and (* The tactic Double performs a double induction *) let simple_elimination c = - elimination_then (fun _ -> tclIDTAC) (EConstr.of_constr c) + elimination_then (fun _ -> tclIDTAC) c let induction_trailer abs_i abs_j bargs = tclTHEN @@ -135,8 +135,9 @@ let induction_trailer abs_i abs_j bargs = (onLastHypId (fun id -> Proofview.Goal.nf_enter { enter = begin fun gl -> - let idty = pf_unsafe_type_of gl (EConstr.mkVar id) in - let fvty = global_vars (pf_env gl) (project gl) (EConstr.of_constr idty) in + let idty = pf_unsafe_type_of gl (mkVar id) in + let idty = EConstr.of_constr idty in + let fvty = global_vars (pf_env gl) (project gl) idty in let possible_bring_hyps = (List.tl (nLastDecls gl (abs_j - abs_i))) @ bargs.Tacticals.assums in @@ -168,7 +169,7 @@ let double_ind h1 h2 = (onLastHypId (fun id -> elimination_then - (introElimAssumsThen (induction_trailer abs_i abs_j)) (EConstr.mkVar id)))) + (introElimAssumsThen (induction_trailer abs_i abs_j)) (mkVar id)))) end } let h_double_induction = double_ind diff --git a/tactics/elim.mli b/tactics/elim.mli index 29c441463..dc1af79ba 100644 --- a/tactics/elim.mli +++ b/tactics/elim.mli @@ -8,6 +8,7 @@ open Names open Term +open EConstr open Tacticals open Misctypes open Tactypes diff --git a/tactics/tacticals.ml b/tactics/tacticals.ml index e15ee149d..e440f1dc5 100644 --- a/tactics/tacticals.ml +++ b/tactics/tacticals.ml @@ -128,7 +128,7 @@ let onClauseLR tac cl gls = tclMAP tac (List.rev (Locusops.simple_clause_of hyps cl)) gls let ifOnHyp pred tac1 tac2 id gl = - if pred (id,pf_get_hyp_typ gl id) then + if pred (id,EConstr.of_constr (pf_get_hyp_typ gl id)) then tac1 id gl else tac2 id gl @@ -583,6 +583,7 @@ module New = struct let ifOnHyp pred tac1 tac2 id = Proofview.Goal.nf_enter { enter = begin fun gl -> let typ = Tacmach.New.pf_get_hyp_typ id gl in + let typ = EConstr.of_constr typ in if pred (id,typ) then tac1 id else diff --git a/tactics/tacticals.mli b/tactics/tacticals.mli index 2c3e51280..e4f110722 100644 --- a/tactics/tacticals.mli +++ b/tactics/tacticals.mli @@ -76,7 +76,7 @@ val nLastDecls : int -> goal sigma -> Context.Named.t val afterHyp : Id.t -> goal sigma -> Context.Named.t -val ifOnHyp : (Id.t * types -> bool) -> +val ifOnHyp : (Id.t * EConstr.types -> bool) -> (Id.t -> tactic) -> (Id.t -> tactic) -> Id.t -> tactic @@ -230,7 +230,7 @@ module New : sig val nLastDecls : ([ `NF ], 'r) Proofview.Goal.t -> int -> Context.Named.t - val ifOnHyp : (identifier * types -> bool) -> + val ifOnHyp : (identifier * EConstr.types -> bool) -> (identifier -> unit Proofview.tactic) -> (identifier -> unit Proofview.tactic) -> identifier -> unit Proofview.tactic -- cgit v1.2.3 From 3f9e56fcbf479999325a86bbdaeefd6a0be13c65 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 18 Nov 2016 20:35:01 +0100 Subject: Equality API using EConstr. --- engine/eConstr.ml | 5 + engine/eConstr.mli | 6 +- engine/evarutil.ml | 17 +- engine/evarutil.mli | 2 +- engine/termops.ml | 2 +- engine/termops.mli | 1 + ltac/extratactics.ml4 | 18 +- ltac/rewrite.ml | 1 + plugins/cc/cctac.ml | 3 +- plugins/fourier/fourierR.ml | 12 +- plugins/funind/functional_principles_proofs.ml | 10 +- plugins/funind/indfun_common.ml | 1 + plugins/funind/invfun.ml | 10 +- plugins/funind/recdef.ml | 12 +- tactics/autorewrite.ml | 4 +- tactics/class_tactics.ml | 4 +- tactics/eqdecide.ml | 2 +- tactics/equality.ml | 304 +++++++++++++------------ tactics/equality.mli | 21 +- tactics/hints.ml | 2 +- tactics/inv.ml | 10 +- toplevel/auto_ind_decl.ml | 12 +- 22 files changed, 241 insertions(+), 218 deletions(-) diff --git a/engine/eConstr.ml b/engine/eConstr.ml index 9e0a55a0d..1dd9d0c00 100644 --- a/engine/eConstr.ml +++ b/engine/eConstr.ml @@ -78,6 +78,8 @@ type cofixpoint = (t, t) pcofixpoint type unsafe_judgment = (constr, types) Environ.punsafe_judgment type unsafe_type_judgment = types Environ.punsafe_type_judgment +let in_punivs a = (a, Univ.Instance.empty) + let mkProp = of_kind (Sort Sorts.prop) let mkSet = of_kind (Sort Sorts.set) let mkType u = of_kind (Sort (Sorts.Type u)) @@ -92,8 +94,11 @@ 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 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) diff --git a/engine/eConstr.mli b/engine/eConstr.mli index 15463a8f6..e6270fa78 100644 --- a/engine/eConstr.mli +++ b/engine/eConstr.mli @@ -62,12 +62,12 @@ 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 mkConst : constant -> t val mkConstU : pconstant -> t val mkProj : (projection * t) -> t -(* val mkInd : inductive -> t *) +val mkInd : inductive -> t val mkIndU : pinductive -> t -(* val mkConstruct : constructor -> t *) +val mkConstruct : constructor -> t val mkConstructU : pconstructor -> t (* val mkConstructUi : pinductive * int -> t *) val mkCase : case_info * t * t * t array -> t diff --git a/engine/evarutil.ml b/engine/evarutil.ml index 4f40499d0..c2ad3c462 100644 --- a/engine/evarutil.ml +++ b/engine/evarutil.ml @@ -690,29 +690,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 = Context.Named.fold_outside - (NamedDecl.fold_constr (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))) diff --git a/engine/evarutil.mli b/engine/evarutil.mli index 6620bbaed..82346b24e 100644 --- a/engine/evarutil.mli +++ b/engine/evarutil.mli @@ -121,7 +121,7 @@ val advance : evar_map -> evar -> evar option This is roughly a combination of the previous functions and [nf_evar]. *) -val undefined_evars_of_term : evar_map -> constr -> Evar.Set.t +val undefined_evars_of_term : evar_map -> EConstr.constr -> Evar.Set.t val undefined_evars_of_named_context : evar_map -> Context.Named.t -> Evar.Set.t val undefined_evars_of_evar_info : evar_map -> evar_info -> Evar.Set.t diff --git a/engine/termops.ml b/engine/termops.ml index b7932665a..c2d862f00 100644 --- a/engine/termops.ml +++ b/engine/termops.ml @@ -1074,7 +1074,7 @@ 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.Unsafe.to_constr c)) acc + Id.Set.union (vars_of_global env (EConstr.to_constr sigma c)) acc | _ -> acc in EConstr.fold sigma filtrec acc c diff --git a/engine/termops.mli b/engine/termops.mli index 7758a57ee..013efcbcb 100644 --- a/engine/termops.mli +++ b/engine/termops.mli @@ -256,6 +256,7 @@ val compact_named_context : Context.Named.t -> Context.Compacted.t val clear_named_body : Id.t -> env -> env val global_vars : env -> Evd.evar_map -> EConstr.t -> Id.t list +val global_vars_set : env -> Evd.evar_map -> EConstr.t -> Id.Set.t val global_vars_set_of_decl : env -> Evd.evar_map -> Context.Named.Declaration.t -> Id.Set.t val global_app_of_constr : Evd.evar_map -> EConstr.constr -> Globnames.global_reference puniverses * EConstr.constr option diff --git a/ltac/extratactics.ml4 b/ltac/extratactics.ml4 index faf545d4f..bcfa13c79 100644 --- a/ltac/extratactics.ml4 +++ b/ltac/extratactics.ml4 @@ -47,16 +47,16 @@ let with_delayed_uconstr ist c tac = let replace_in_clause_maybe_by ist c1 c2 cl tac = with_delayed_uconstr ist c1 - (fun c1 -> replace_in_clause_maybe_by c1 c2 cl (Option.map (Tacinterp.tactic_of_value ist) tac)) + (fun c1 -> replace_in_clause_maybe_by (EConstr.of_constr c1) c2 cl (Option.map (Tacinterp.tactic_of_value ist) tac)) let replace_term ist dir_opt c cl = - with_delayed_uconstr ist c (fun c -> replace_term dir_opt c cl) + with_delayed_uconstr ist c (fun c -> replace_term dir_opt (EConstr.of_constr c) cl) let clause = Pltac.clause_dft_concl TACTIC EXTEND replace ["replace" uconstr(c1) "with" constr(c2) clause(cl) by_arg_tac(tac) ] --> [ replace_in_clause_maybe_by ist c1 c2 cl tac ] +-> [ replace_in_clause_maybe_by ist c1 (EConstr.of_constr c2) cl tac ] END TACTIC EXTEND replace_term_left @@ -153,9 +153,9 @@ let injHyp id = injection_main false { delayed = fun env sigma -> Sigma.here (EConstr.mkVar id, NoBindings) sigma } TACTIC EXTEND dependent_rewrite -| [ "dependent" "rewrite" orient(b) constr(c) ] -> [ rewriteInConcl b c ] +| [ "dependent" "rewrite" orient(b) constr(c) ] -> [ rewriteInConcl b (EConstr.of_constr c) ] | [ "dependent" "rewrite" orient(b) constr(c) "in" hyp(id) ] - -> [ rewriteInHyp b c id ] + -> [ rewriteInHyp b (EConstr.of_constr c) id ] END (** To be deprecated?, "cutrewrite (t=u) as <-" is equivalent to @@ -163,9 +163,9 @@ END "cutrewrite (t=u) as ->" is equivalent to "enough (t=u) as ->". *) TACTIC EXTEND cut_rewrite -| [ "cutrewrite" orient(b) constr(eqn) ] -> [ cutRewriteInConcl b eqn ] +| [ "cutrewrite" orient(b) constr(eqn) ] -> [ cutRewriteInConcl b (EConstr.of_constr eqn) ] | [ "cutrewrite" orient(b) constr(eqn) "in" hyp(id) ] - -> [ cutRewriteInHyp b eqn id ] + -> [ cutRewriteInHyp b (EConstr.of_constr eqn) id ] END (**********************************************************************) @@ -235,7 +235,7 @@ END let rewrite_star ist clause orient occs c (tac : Geninterp.Val.t option) = let tac' = Option.map (fun t -> Tacinterp.tactic_of_value ist t, FirstSolved) tac in with_delayed_uconstr ist c - (fun c -> general_rewrite_ebindings_clause clause orient occs ?tac:tac' true true (c,NoBindings) true) + (fun c -> general_rewrite_ebindings_clause clause orient occs ?tac:tac' true true (EConstr.of_constr c,NoBindings) true) TACTIC EXTEND rewrite_star | [ "rewrite" "*" orient(o) uconstr(c) "in" hyp(id) "at" occurrences(occ) by_arg_tac(tac) ] -> @@ -719,7 +719,7 @@ let rewrite_except h = Proofview.Goal.nf_enter { enter = begin fun gl -> let hyps = Tacmach.New.pf_ids_of_hyps gl in Tacticals.New.tclMAP (fun id -> if Id.equal id h then Proofview.tclUNIT () else - Tacticals.New.tclTRY (Equality.general_rewrite_in true Locus.AllOccurrences true true id (mkVar h) false)) + Tacticals.New.tclTRY (Equality.general_rewrite_in true Locus.AllOccurrences true true id (EConstr.mkVar h) false)) hyps end } diff --git a/ltac/rewrite.ml b/ltac/rewrite.ml index ef2ab0917..0d279ae93 100644 --- a/ltac/rewrite.ml +++ b/ltac/rewrite.ml @@ -2094,6 +2094,7 @@ let general_rewrite_flags = { under_lambdas = false; on_morphisms = true } (** Setoid rewriting when called with "rewrite" *) let general_s_rewrite cl l2r occs (c,l) ~new_goals = Proofview.Goal.nf_enter { enter = begin fun gl -> + let (c,l) = Miscops.map_with_bindings EConstr.Unsafe.to_constr (c,l) in let abs, evd, res, sort = get_hyp gl (c,l) cl l2r in let unify env evars t = unify_abs res l2r sort env evars t in let app = apply_rule unify occs in diff --git a/plugins/cc/cctac.ml b/plugins/cc/cctac.ml index 7b023413d..a12ef00ec 100644 --- a/plugins/cc/cctac.ml +++ b/plugins/cc/cctac.ml @@ -229,7 +229,8 @@ let make_prb gls depth additionnal_terms = let build_projection intype (cstr:pconstructor) special default gls= let ci= (snd(fst cstr)) in - let body=Equality.build_selector (pf_env gls) (project gls) ci (mkRel 1) intype special default in + let body=Equality.build_selector (pf_env gls) (project gls) ci (EConstr.mkRel 1) (EConstr.of_constr intype) (EConstr.of_constr special) (EConstr.of_constr default) in + let body = EConstr.Unsafe.to_constr body in let id=pf_get_new_id (Id.of_string "t") gls in mkLambda(Name id,intype,body) diff --git a/plugins/fourier/fourierR.ml b/plugins/fourier/fourierR.ml index fa64b276c..dbb5cc2de 100644 --- a/plugins/fourier/fourierR.ml +++ b/plugins/fourier/fourierR.ml @@ -600,15 +600,15 @@ let rec fourier () = (Tacticals.New.tclTHEN (apply (if sres then get coq_Rnot_lt_lt else get coq_Rnot_le_le)) (Tacticals.New.tclTHENS (Equality.replace - (mkAppL [|cget coq_Rminus;!t2;!t1|] - ) - tc) + (EConstr.of_constr (mkAppL [|cget coq_Rminus;!t2;!t1|] + )) + (EConstr.of_constr tc)) [tac2; (Tacticals.New.tclTHENS (Equality.replace - (mkApp (cget coq_Rinv, - [|cget coq_R1|])) - (cget coq_R1)) + (EConstr.of_constr (mkApp (cget coq_Rinv, + [|cget coq_R1|]))) + (get coq_R1)) (* en attendant Field, ça peut aider Ring de remplacer 1/1 par 1 ... *) [Tacticals.New.tclORELSE diff --git a/plugins/funind/functional_principles_proofs.ml b/plugins/funind/functional_principles_proofs.ml index 940f1669a..2e3992be9 100644 --- a/plugins/funind/functional_principles_proofs.ml +++ b/plugins/funind/functional_principles_proofs.ml @@ -401,7 +401,7 @@ let rewrite_until_var arg_num eq_ids : tactic = | [] -> anomaly (Pp.str "Cannot find a way to prove recursive property"); | eq_id::eq_ids -> tclTHEN - (tclTRY (Proofview.V82.of_tactic (Equality.rewriteRL (mkVar eq_id)))) + (tclTRY (Proofview.V82.of_tactic (Equality.rewriteRL (EConstr.mkVar eq_id)))) (do_rewrite eq_ids) g in @@ -1060,7 +1060,7 @@ let do_replace (evd:Evd.evar_map ref) params rec_arg_num rev_args_id f fun_num a let just_introduced = nLastDecls nb_intro_to_do g' in let open Context.Named.Declaration in let just_introduced_id = List.map get_id just_introduced in - tclTHEN (Proofview.V82.of_tactic (Equality.rewriteLR equation_lemma)) + tclTHEN (Proofview.V82.of_tactic (Equality.rewriteLR (EConstr.of_constr equation_lemma))) (revert just_introduced_id) g' ) g @@ -1425,7 +1425,7 @@ let prove_with_tcc tcc_lemma_constr eqs : tactic = let backtrack_eqs_until_hrec hrec eqs : tactic = fun gls -> - let eqs = List.map mkVar eqs in + let eqs = List.map EConstr.mkVar eqs in let rewrite = tclFIRST (List.map (fun x -> Proofview.V82.of_tactic (Equality.rewriteRL x)) eqs ) in @@ -1453,7 +1453,7 @@ let rec rewrite_eqs_in_eqs eqs = observe_tac (Format.sprintf "rewrite %s in %s " (Id.to_string eq) (Id.to_string id)) (tclTRY (Proofview.V82.of_tactic (Equality.general_rewrite_in true Locus.AllOccurrences - true (* dep proofs also: *) true id (mkVar eq) false))) + true (* dep proofs also: *) true id (EConstr.mkVar eq) false))) gl ) eqs @@ -1659,7 +1659,7 @@ let prove_principle_for_gen (* observe_tac "h_fix " *) (Proofview.V82.of_tactic (fix (Some fix_id) (List.length args_ids + 1))); (* (fun g -> observe (Printer.pr_goal (sig_it g) ++ fnl() ++ pr_lconstr_env (pf_env g ) (pf_unsafe_type_of g (mkVar fix_id) )); tclIDTAC g); *) h_intros (List.rev (acc_rec_arg_id::args_ids)); - Proofview.V82.of_tactic (Equality.rewriteLR (mkConst eq_ref)); + Proofview.V82.of_tactic (Equality.rewriteLR (EConstr.of_constr (mkConst eq_ref))); (* observe_tac "finish" *) (fun gl' -> let body = let _,args = destApp (pf_concl gl') in diff --git a/plugins/funind/indfun_common.ml b/plugins/funind/indfun_common.ml index a45effb16..08b40a1f7 100644 --- a/plugins/funind/indfun_common.ml +++ b/plugins/funind/indfun_common.ml @@ -502,6 +502,7 @@ let evaluable_of_global_reference r = (* Tacred.evaluable_of_global_reference (G | _ -> assert false;; let list_rewrite (rev:bool) (eqs: (constr*bool) list) = + let eqs = List.map (Util.on_fst EConstr.of_constr) eqs in tclREPEAT (List.fold_right (fun (eq,b) i -> tclORELSE (Proofview.V82.of_tactic ((if b then Equality.rewriteLR else Equality.rewriteRL) eq)) i) diff --git a/plugins/funind/invfun.ml b/plugins/funind/invfun.ml index 36fb6dad3..d29d4694f 100644 --- a/plugins/funind/invfun.ml +++ b/plugins/funind/invfun.ml @@ -385,7 +385,7 @@ let prove_fun_correct evd functional_induction funs_constr graphs_constr schemes (* introducing the the result of the graph and the equality hypothesis *) observe_tac "introducing" (tclMAP (fun x -> Proofview.V82.of_tactic (Simple.intro x)) [res;hres]); (* replacing [res] with its value *) - observe_tac "rewriting res value" (Proofview.V82.of_tactic (Equality.rewriteLR (mkVar hres))); + observe_tac "rewriting res value" (Proofview.V82.of_tactic (Equality.rewriteLR (EConstr.mkVar hres))); (* Conclusion *) observe_tac "exact" (fun g -> Proofview.V82.of_tactic (exact_check (EConstr.of_constr (app_constructor g))) g) @@ -520,7 +520,7 @@ and intros_with_rewrite_aux : tactic = let id = pf_get_new_id (Id.of_string "y") g in tclTHENSEQ [ Proofview.V82.of_tactic (Simple.intro id); generalize_dependent_of (destVar args.(1)) id; - tclTRY (Proofview.V82.of_tactic (Equality.rewriteLR (mkVar id))); + tclTRY (Proofview.V82.of_tactic (Equality.rewriteLR (EConstr.mkVar id))); intros_with_rewrite ] g @@ -529,7 +529,7 @@ and intros_with_rewrite_aux : tactic = let id = pf_get_new_id (Id.of_string "y") g in tclTHENSEQ [ Proofview.V82.of_tactic (Simple.intro id); generalize_dependent_of (destVar args.(2)) id; - tclTRY (Proofview.V82.of_tactic (Equality.rewriteRL (mkVar id))); + tclTRY (Proofview.V82.of_tactic (Equality.rewriteRL (EConstr.mkVar id))); intros_with_rewrite ] g @@ -538,7 +538,7 @@ and intros_with_rewrite_aux : tactic = let id = pf_get_new_id (Id.of_string "y") g in tclTHENSEQ[ Proofview.V82.of_tactic (Simple.intro id); - tclTRY (Proofview.V82.of_tactic (Equality.rewriteLR (mkVar id))); + tclTRY (Proofview.V82.of_tactic (Equality.rewriteLR (EConstr.mkVar id))); intros_with_rewrite ] g end @@ -709,7 +709,7 @@ let prove_fun_complete funcs graphs schemes lemmas_types_infos i : tactic = in tclTHENSEQ[ tclMAP (fun id -> Proofview.V82.of_tactic (Simple.intro id)) ids; - Proofview.V82.of_tactic (Equality.rewriteLR (mkConst eq_lemma)); + Proofview.V82.of_tactic (Equality.rewriteLR (EConstr.of_constr (mkConst eq_lemma))); (* Don't forget to $\zeta$ normlize the term since the principles have been $\zeta$-normalized *) Proofview.V82.of_tactic (reduce diff --git a/plugins/funind/recdef.ml b/plugins/funind/recdef.ml index 74affa396..5cee3cb20 100644 --- a/plugins/funind/recdef.ml +++ b/plugins/funind/recdef.ml @@ -871,10 +871,10 @@ let rec make_rewrite_list expr_info max = function in Proofview.V82.of_tactic (general_rewrite_bindings false Locus.AllOccurrences true (* dep proofs also: *) true - (mkVar hp, + (EConstr.mkVar hp, ExplicitBindings[Loc.ghost,NamedHyp def, - expr_info.f_constr;Loc.ghost,NamedHyp k, - (f_S max)]) false) g) ) + EConstr.of_constr expr_info.f_constr;Loc.ghost,NamedHyp k, + EConstr.of_constr (f_S max)]) false) g) ) ) [make_rewrite_list expr_info max l; observe_tclTHENLIST (str "make_rewrite_list")[ (* x < S max proof *) @@ -898,10 +898,10 @@ let make_rewrite expr_info l hp max = observe_tac (str "general_rewrite_bindings") (Proofview.V82.of_tactic (general_rewrite_bindings false Locus.AllOccurrences true (* dep proofs also: *) true - (mkVar hp, + (EConstr.mkVar hp, ExplicitBindings[Loc.ghost,NamedHyp def, - expr_info.f_constr;Loc.ghost,NamedHyp k, - (f_S (f_S max))]) false)) g) + EConstr.of_constr expr_info.f_constr;Loc.ghost,NamedHyp k, + EConstr.of_constr (f_S (f_S max))]) false)) g) [observe_tac(str "make_rewrite finalize") ( (* tclORELSE( h_reflexivity) *) (observe_tclTHENLIST (str "make_rewrite")[ diff --git a/tactics/autorewrite.ml b/tactics/autorewrite.ml index b567344c9..d656206d6 100644 --- a/tactics/autorewrite.ml +++ b/tactics/autorewrite.ml @@ -122,7 +122,7 @@ let autorewrite ?(conds=Naive) tac_main lbas = Tacticals.New.tclTHEN tac (one_base (fun dir c tac -> let tac = (tac, conds) in - general_rewrite dir AllOccurrences true false ~tac c) + general_rewrite dir AllOccurrences true false ~tac (EConstr.of_constr c)) tac_main bas)) (Proofview.tclUNIT()) lbas)) @@ -165,7 +165,7 @@ let autorewrite_multi_in ?(conds=Naive) idl tac_main lbas = | _ -> assert false) (* there must be at least an hypothesis *) | _ -> assert false (* rewriting cannot complete a proof *) in - let general_rewrite_in x y z w = Proofview.V82.tactic (general_rewrite_in x y z w) in + let general_rewrite_in x y z w = Proofview.V82.tactic (general_rewrite_in x y (EConstr.of_constr z) w) in Tacticals.New.tclMAP (fun id -> Tacticals.New.tclREPEAT_MAIN (Proofview.tclPROGRESS (List.fold_left (fun tac bas -> diff --git a/tactics/class_tactics.ml b/tactics/class_tactics.ml index 7d8fc79f4..02211efd6 100644 --- a/tactics/class_tactics.ml +++ b/tactics/class_tactics.ml @@ -1315,8 +1315,8 @@ module Intpart = Unionfind.Make(Evar.Set)(Evar.Map) let deps_of_constraints cstrs evm p = List.iter (fun (_, _, x, y) -> - let evx = Evarutil.undefined_evars_of_term evm x in - let evy = Evarutil.undefined_evars_of_term evm y in + let evx = Evarutil.undefined_evars_of_term evm (EConstr.of_constr x) in + let evy = Evarutil.undefined_evars_of_term evm (EConstr.of_constr y) in Intpart.union_set (Evar.Set.union evx evy) p) cstrs diff --git a/tactics/eqdecide.ml b/tactics/eqdecide.ml index eb75cbf7d..be9a34239 100644 --- a/tactics/eqdecide.ml +++ b/tactics/eqdecide.ml @@ -116,7 +116,7 @@ let rec rewrite_and_clear hyps = match hyps with | [] -> Proofview.tclUNIT () | id :: hyps -> tclTHENLIST [ - Equality.rewriteLR (mkVar id); + Equality.rewriteLR (EConstr.mkVar id); clear [id]; rewrite_and_clear hyps; ] diff --git a/tactics/equality.ml b/tactics/equality.ml index 80f83f19b..4c79a6199 100644 --- a/tactics/equality.ml +++ b/tactics/equality.ml @@ -12,8 +12,9 @@ open Util open Names open Nameops open Term -open Vars open Termops +open EConstr +open Vars open Namegen open Inductive open Inductiveops @@ -46,6 +47,10 @@ open Context.Named.Declaration module NamedDecl = Context.Named.Declaration +let nlocal_assum (na, t) = + let inj = EConstr.Unsafe.to_constr in + NamedDecl.LocalAssum (na, inj t) + (* Options *) let discriminate_introduction = ref true @@ -144,7 +149,7 @@ let freeze_initial_evars sigma flags clause = (* We take evars of the type: this may include old evars! For excluding *) (* all old evars, including the ones occurring in the rewriting lemma, *) (* we would have to take the clenv_value *) - let newevars = Evd.evars_of_term (EConstr.Unsafe.to_constr (clenv_type clause)) in + let newevars = Evarutil.undefined_evars_of_term sigma (clenv_type clause) in let evars = fold_undefined (fun evk _ evars -> if Evar.Set.mem evk newevars then evars @@ -165,11 +170,9 @@ let side_tac tac sidetac = let instantiate_lemma_all frzevars gl c ty l l2r concl = let env = Proofview.Goal.env gl in - let c = EConstr.of_constr c in - let ty = EConstr.of_constr ty in - let l = Miscops.map_bindings EConstr.of_constr l in + let sigma = project gl in let eqclause = pf_apply Clenv.make_clenv_binding gl (c,ty) l in - let (equiv, args) = decompose_appvect (EConstr.Unsafe.to_constr (Clenv.clenv_type eqclause)) in + let (equiv, args) = decompose_app_vect sigma (Clenv.clenv_type eqclause) in let arglen = Array.length args in let () = if arglen < 2 then error "The term provided is not an applied relation." in let c1 = args.(arglen - 2) in @@ -184,11 +187,9 @@ let instantiate_lemma_all frzevars gl c ty l l2r concl = in List.map try_occ occs let instantiate_lemma gl c ty l l2r concl = - let c = EConstr.of_constr c in let sigma, ct = pf_type_of gl c in let ct = EConstr.of_constr ct in let t = try snd (reduce_to_quantified_ind (pf_env gl) sigma ct) with UserError _ -> ct in - let l = Miscops.map_bindings EConstr.of_constr l in let eqclause = Clenv.make_clenv_binding (pf_env gl) sigma (c,t) l in [eqclause] @@ -332,9 +333,9 @@ let (forward_general_setoid_rewrite_clause, general_setoid_rewrite_clause) = Hoo let jmeq_same_dom gl = function | None -> true (* already checked in Hipattern.find_eq_data_decompose *) | Some t -> - let rels, t = decompose_prod_assum t in + let rels, t = decompose_prod_assum (project gl) t in let env = Environ.push_rel_context rels (Proofview.Goal.env gl) in - match EConstr.decompose_app (project gl) (EConstr.of_constr t) with + match decompose_app (project gl) t with | _, [dom1; _; dom2;_] -> is_conv env (Tacmach.New.project gl) dom1 dom2 | _ -> false @@ -342,6 +343,8 @@ let jmeq_same_dom gl = function eliminate lbeq on sort_of_gl. *) let find_elim hdcncl lft2rgt dep cls ot gl = + let sigma = project gl in + let is_global gr c = Termops.is_global sigma gr c in let inccl = Option.is_empty cls in if (is_global Coqlib.glob_eq hdcncl || (is_global Coqlib.glob_jmeq hdcncl && @@ -349,7 +352,7 @@ let find_elim hdcncl lft2rgt dep cls ot gl = || Flags.version_less_or_equal Flags.V8_2 then let c = - match kind_of_term hdcncl with + match EConstr.kind sigma hdcncl with | Ind (ind_sp,u) -> let pr1 = lookup_eliminator ind_sp (elimination_sort_of_clause cls gl) @@ -377,6 +380,7 @@ let find_elim hdcncl lft2rgt dep cls ot gl = assert false in let Sigma (elim, sigma, p) = Sigma.fresh_global (Global.env ()) (Proofview.Goal.sigma gl) (ConstRef c) in + let elim = EConstr.of_constr elim in Sigma ((elim, Safe_typing.empty_private_constants), sigma, p) else let scheme_name = match dep, lft2rgt, inccl with @@ -391,13 +395,14 @@ let find_elim hdcncl lft2rgt dep cls ot gl = | true, _, true -> rew_r2l_dep_scheme_kind | true, _, false -> rew_r2l_forward_dep_scheme_kind in - match kind_of_term hdcncl with + match EConstr.kind sigma hdcncl with | Ind (ind,u) -> let c, eff = find_scheme scheme_name ind in (* MS: cannot use pf_constr_of_global as the eliminator might be generated by side-effect *) let Sigma (elim, sigma, p) = Sigma.fresh_global (Global.env ()) (Proofview.Goal.sigma gl) (ConstRef c) in + let elim = EConstr.of_constr elim in Sigma ((elim, eff), sigma, p) | _ -> assert false @@ -408,12 +413,12 @@ let type_of_clause cls gl = match cls with let leibniz_rewrite_ebindings_clause cls lft2rgt tac c t l with_evars frzevars dep_proof_ok hdcncl = Proofview.Goal.nf_s_enter { s_enter = begin fun gl -> let evd = Sigma.to_evar_map (Proofview.Goal.sigma gl) in - let isatomic = isProd (whd_zeta evd (EConstr.of_constr hdcncl)) in + let isatomic = isProd evd (EConstr.of_constr (whd_zeta evd hdcncl)) in let dep_fun = if isatomic then dependent else dependent_no_evar in let type_of_cls = type_of_clause cls gl in - let dep = dep_proof_ok && dep_fun evd (EConstr.of_constr c) (EConstr.of_constr type_of_cls) in + let type_of_cls = EConstr.of_constr type_of_cls in + let dep = dep_proof_ok && dep_fun evd c type_of_cls in let Sigma ((elim, effs), sigma, p) = find_elim hdcncl lft2rgt dep cls (Some t) gl in - let elim = EConstr.of_constr elim in let tac = Proofview.tclEFFECTS effs <*> general_elim_clause with_evars frzevars tac cls c t l @@ -447,11 +452,11 @@ let general_rewrite_ebindings_clause cls lft2rgt occs frzevars dep_proof_ok ?tac Proofview.Goal.enter { enter = begin fun gl -> let sigma = Tacmach.New.project gl in let env = Proofview.Goal.env gl in - let ctype = get_type_of env sigma (EConstr.of_constr c) in - let rels, t = decompose_prod_assum (whd_betaiotazeta sigma (EConstr.of_constr ctype)) in - match match_with_equality_type sigma (EConstr.of_constr t) with + let ctype = get_type_of env sigma c in + let ctype = EConstr.of_constr ctype in + let rels, t = decompose_prod_assum sigma (EConstr.of_constr (whd_betaiotazeta sigma ctype)) in + match match_with_equality_type sigma t with | Some (hdcncl,args) -> (* Fast path: direct leibniz-like rewrite *) - let hdcncl = EConstr.Unsafe.to_constr hdcncl in let lft2rgt = adjust_rewriting_direction args lft2rgt in leibniz_rewrite_ebindings_clause cls lft2rgt tac c (it_mkProd_or_LetIn t rels) l with_evars frzevars dep_proof_ok hdcncl @@ -465,10 +470,10 @@ let general_rewrite_ebindings_clause cls lft2rgt occs frzevars dep_proof_ok ?tac | (e, info) -> Proofview.tclEVARMAP >>= fun sigma -> let env' = push_rel_context rels env in - let rels',t' = splay_prod_assum env' sigma (EConstr.of_constr t) in (* Search for underlying eq *) - match match_with_equality_type sigma (EConstr.of_constr t') with + let rels',t' = splay_prod_assum env' sigma t in (* Search for underlying eq *) + let t' = EConstr.of_constr t' in + match match_with_equality_type sigma t' with | Some (hdcncl,args) -> - let hdcncl = EConstr.Unsafe.to_constr hdcncl in let lft2rgt = adjust_rewriting_direction args lft2rgt in leibniz_rewrite_ebindings_clause cls lft2rgt tac c (it_mkProd_or_LetIn t' (rels' @ rels)) l with_evars frzevars dep_proof_ok hdcncl @@ -533,7 +538,7 @@ let general_rewrite_clause l2r with_evars ?tac c cl = let do_hyps = (* If the term to rewrite uses an hypothesis H, don't rewrite in H *) let ids gl = - let ids_in_c = Environ.global_vars_set (Global.env()) (fst c) in + let ids_in_c = Termops.global_vars_set (Global.env()) (project gl) (fst c) in let ids_of_hyps = pf_ids_of_hyps gl in Id.Set.fold (fun id l -> List.remove Id.equal id l) ids_in_c ids_of_hyps in @@ -563,7 +568,6 @@ let general_multi_rewrite with_evars l cl tac = let sigma = Tacmach.New.project gl in let env = Proofview.Goal.env gl in let (c, sigma) = run_delayed env sigma f in - let c = Miscops.map_with_bindings EConstr.Unsafe.to_constr c in tclWITHHOLES with_evars (general_rewrite_clause l2r with_evars ?tac c cl) sigma end } @@ -631,12 +635,14 @@ let replace_using_leibniz clause c1 c2 l2r unsafe try_prove_eq_opt = in Proofview.Goal.enter { enter = begin fun gl -> let get_type_of = pf_apply get_type_of gl in - let t1 = get_type_of (EConstr.of_constr c1) - and t2 = get_type_of (EConstr.of_constr c2) in + let t1 = get_type_of c1 + and t2 = get_type_of c2 in + let t1 = EConstr.of_constr t1 in + let t2 = EConstr.of_constr t2 in let evd = if unsafe then Some (Tacmach.New.project gl) else - try Some (Evarconv.the_conv_x (Proofview.Goal.env gl) (EConstr.of_constr t1) (EConstr.of_constr t2) (Tacmach.New.project gl)) + try Some (Evarconv.the_conv_x (Proofview.Goal.env gl) t1 t2 (Tacmach.New.project gl)) with Evarconv.UnableToUnify _ -> None in match evd with @@ -647,9 +653,9 @@ let replace_using_leibniz clause c1 c2 l2r unsafe try_prove_eq_opt = let sym = build_coq_eq_sym () in Tacticals.New.pf_constr_of_global sym (fun sym -> Tacticals.New.pf_constr_of_global e (fun e -> + let e = EConstr.of_constr e in let eq = applist (e, [t1;c1;c2]) in let sym = EConstr.of_constr sym in - let eq = EConstr.of_constr eq in tclTHENLAST (replace_core clause l2r eq) (tclFIRST @@ -727,12 +733,10 @@ let _ = optwrite = (fun b -> keep_proof_equalities_for_injection := b) } let find_positions env sigma t1 t2 = - let open EConstr in let project env sorts posn t1 t2 = - let t1 = EConstr.Unsafe.to_constr t1 in - let t2 = EConstr.Unsafe.to_constr t2 in - let ty1 = get_type_of env sigma (EConstr.of_constr t1) in - let s = get_sort_family_of env sigma (EConstr.of_constr ty1) in + let ty1 = get_type_of env sigma t1 in + let ty1 = EConstr.of_constr ty1 in + let s = get_sort_family_of env sigma ty1 in if Sorts.List.mem s sorts then [(List.rev posn,t1,t2)] else [] in @@ -854,7 +858,7 @@ let injectable env sigma t1 t2 = let descend_then env sigma head dirn = let IndType (indf,_) = - try find_rectype env sigma (EConstr.of_constr (get_type_of env sigma (EConstr.of_constr head))) + try find_rectype env sigma (EConstr.of_constr (get_type_of env sigma head)) with Not_found -> error "Cannot project on an inductive type derived from a dependency." in let indp,_ = (dest_ind_family indf) in @@ -871,12 +875,12 @@ let descend_then env sigma head dirn = it_mkLambda_or_LetIn (lift (mip.mind_nrealargs+1) resty) deparsign in let build_branch i = let result = if Int.equal i dirn then dirnval else dfltval in - it_mkLambda_or_LetIn_name env result cstr.(i-1).cs_args in + it_mkLambda_or_LetIn result (name_context env cstr.(i-1).cs_args) in let brl = List.map build_branch (List.interval 1 (Array.length mip.mind_consnames)) in let ci = make_case_info env ind RegularStyle in - EConstr.Unsafe.to_constr (Inductiveops.make_case_or_project env sigma indf ci (EConstr.of_constr p) (EConstr.of_constr head) (Array.map_of_list EConstr.of_constr brl)))) + Inductiveops.make_case_or_project env sigma indf ci p head (Array.of_list brl))) (* Now we need to construct the discriminator, given a discriminable position. This boils down to: @@ -897,7 +901,7 @@ let descend_then env sigma head dirn = let build_selector env sigma dirn c ind special default = let IndType(indf,_) = - try find_rectype env sigma (EConstr.of_constr ind) + try find_rectype env sigma ind with Not_found -> (* one can find Rel(k) in case of dependent constructors like T := c : (A:Set)A->T and a discrimination @@ -909,7 +913,8 @@ let build_selector env sigma dirn c ind special default = dependent types.") in let (indp,_) = dest_ind_family indf in let ind, _ = check_privacy env indp in - let typ = Retyping.get_type_of env sigma (EConstr.of_constr default) in + let typ = Retyping.get_type_of env sigma default in + let typ = EConstr.of_constr typ in let (mib,mip) = lookup_mind_specif env ind in let deparsign = make_arity_signature env true indf in let p = it_mkLambda_or_LetIn typ deparsign in @@ -922,9 +927,14 @@ let build_selector env sigma dirn c ind special default = let ci = make_case_info env ind RegularStyle in mkCase (ci, p, c, Array.of_list brl) +let build_coq_False () = EConstr.of_constr (build_coq_False ()) +let build_coq_True () = EConstr.of_constr (build_coq_True ()) +let build_coq_I () = EConstr.of_constr (build_coq_I ()) + let rec build_discriminator env sigma dirn c = function | [] -> - let ind = get_type_of env sigma (EConstr.of_constr c) in + let ind = get_type_of env sigma c in + let ind = EConstr.of_constr ind in let true_0,false_0 = build_coq_True(),build_coq_False() in build_selector env sigma dirn c ind true_0 false_0 @@ -952,7 +962,7 @@ let gen_absurdity id = let hyp_typ = EConstr.of_constr hyp_typ in if is_empty_type sigma hyp_typ then - simplest_elim (EConstr.mkVar id) + simplest_elim (mkVar id) else tclZEROMSG (str "Not the negation of an equality.") end } @@ -980,6 +990,7 @@ let discrimination_pf env sigma e (t,t1,t2) discriminator lbeq = let absurd_term = build_coq_False () in let eq_elim, eff = ind_scheme_of_eq lbeq in let sigma, eq_elim = Evd.fresh_global (Global.env ()) sigma eq_elim in + let eq_elim = EConstr.of_constr eq_elim in sigma, (applist (eq_elim, [t;t1;mkNamedLambda e t discriminator;i;t2]), absurd_term), eff @@ -987,8 +998,6 @@ let eq_baseid = Id.of_string "e" let apply_on_clause (f,t) clause = let sigma = clause.evd in - let f = EConstr.of_constr f in - let t = EConstr.of_constr t in let f_clause = mk_clenv_from_env clause.env sigma None (f,t) in let argmv = (match EConstr.kind sigma (last_arg f_clause.evd f_clause.templval.Evd.rebus) with @@ -997,19 +1006,14 @@ let apply_on_clause (f,t) clause = clenv_fchain ~with_univs:false argmv f_clause clause let discr_positions env sigma (lbeq,eqn,(t,t1,t2)) eq_clause cpath dirn = - let t = EConstr.Unsafe.to_constr t in - let t1 = EConstr.Unsafe.to_constr t1 in - let t2 = EConstr.Unsafe.to_constr t2 in - let eqn = EConstr.Unsafe.to_constr eqn in let e = next_ident_away eq_baseid (ids_of_context env) in - let e_env = push_named (Context.Named.Declaration.LocalAssum (e,t)) env in + let e_env = push_named (nlocal_assum (e, t)) env in let discriminator = build_discriminator e_env sigma dirn (mkVar e) cpath in let sigma,(pf, absurd_term), eff = discrimination_pf env sigma e (t,t1,t2) discriminator lbeq in let pf_ty = mkArrow eqn absurd_term in let absurd_clause = apply_on_clause (pf,pf_ty) eq_clause in - let absurd_term = EConstr.of_constr absurd_term in let pf = Clenvtac.clenv_value_cast_meta absurd_clause in Proofview.Unsafe.tclEVARS sigma <*> Proofview.tclEFFECTS eff <*> @@ -1047,19 +1051,20 @@ let onNegatedEquality with_evars tac = Proofview.Goal.nf_enter { enter = begin fun gl -> let sigma = Tacmach.New.project gl in let ccl = Proofview.Goal.concl gl in + let ccl = EConstr.of_constr ccl in let env = Proofview.Goal.env gl in - match kind_of_term (hnf_constr env sigma (EConstr.of_constr ccl)) with - | Prod (_,t,u) when is_empty_type sigma (EConstr.of_constr u) -> + match EConstr.kind sigma (EConstr.of_constr (hnf_constr env sigma ccl)) with + | Prod (_,t,u) when is_empty_type sigma u -> tclTHEN introf (onLastHypId (fun id -> - onEquality with_evars tac (EConstr.mkVar id,NoBindings))) + onEquality with_evars tac (mkVar id,NoBindings))) | _ -> tclZEROMSG (str "Not a negated primitive equality.") end } let discrSimpleClause with_evars = function | None -> onNegatedEquality with_evars discrEq - | Some id -> onEquality with_evars discrEq (EConstr.mkVar id,NoBindings) + | Some id -> onEquality with_evars discrEq (mkVar id,NoBindings) let discr with_evars = onEquality with_evars discrEq @@ -1073,7 +1078,7 @@ let discrEverywhere with_evars = (tclTHEN (tclREPEAT introf) (tryAllHyps - (fun id -> tclCOMPLETE (discr with_evars (EConstr.mkVar id,NoBindings))))) + (fun id -> tclCOMPLETE (discr with_evars (mkVar id,NoBindings))))) else (* <= 8.2 compat *) tryAllHypsAndConcl (discrSimpleClause with_evars)) (* (fun gls -> @@ -1103,9 +1108,10 @@ let find_sigma_data env s = build_sigma_type () *) let make_tuple env sigma (rterm,rty) lind = - assert (not (EConstr.Vars.noccurn sigma lind (EConstr.of_constr rty))); - let sigdata = find_sigma_data env (get_sort_of env sigma (EConstr.of_constr rty)) in - let sigma, a = type_of ~refresh:true env sigma (EConstr.mkRel lind) in + assert (not (noccurn sigma lind rty)); + let sigdata = find_sigma_data env (get_sort_of env sigma rty) in + let sigma, a = type_of ~refresh:true env sigma (mkRel lind) in + let a = EConstr.of_constr a in let na = Context.Rel.Declaration.get_name (lookup_rel lind env) in (* We move [lind] to [1] and lift other rels > [lind] by 1 *) let rty = lift (1-lind) (liftn lind (lind+1) rty) in @@ -1113,6 +1119,8 @@ let make_tuple env sigma (rterm,rty) lind = let p = mkLambda (na, a, rty) in let sigma, exist_term = Evd.fresh_global env sigma sigdata.intro in let sigma, sig_term = Evd.fresh_global env sigma sigdata.typ in + let exist_term = EConstr.of_constr exist_term in + let sig_term = EConstr.of_constr sig_term in sigma, (applist(exist_term,[a;p;(mkRel lind);rterm]), applist(sig_term,[a;p])) @@ -1125,9 +1133,10 @@ let make_tuple env sigma (rterm,rty) lind = normalization *) let minimal_free_rels env sigma (c,cty) = - let cty_rels = free_rels sigma (EConstr.of_constr cty) in - let cty' = simpl env sigma (EConstr.of_constr cty) in - let rels' = free_rels sigma (EConstr.of_constr cty') in + let cty_rels = free_rels sigma cty in + let cty' = simpl env sigma cty in + let cty' = EConstr.of_constr cty' in + let rels' = free_rels sigma cty' in if Int.Set.subset cty_rels rels' then (cty,cty_rels) else @@ -1139,7 +1148,7 @@ let minimal_free_rels_rec env sigma = let rec minimalrec_free_rels_rec prev_rels (c,cty) = let (cty,direct_rels) = minimal_free_rels env sigma (c,cty) in let combined_rels = Int.Set.union prev_rels direct_rels in - let folder rels i = snd (minimalrec_free_rels_rec rels (c, unsafe_type_of env sigma (EConstr.mkRel i))) + let folder rels i = snd (minimalrec_free_rels_rec rels (c, EConstr.of_constr (unsafe_type_of env sigma (mkRel i)))) in (cty, List.fold_left folder combined_rels (Int.Set.elements (Int.Set.diff direct_rels prev_rels))) in minimalrec_free_rels_rec Int.Set.empty @@ -1185,27 +1194,30 @@ let sig_clausal_form env sigma sort_of_ty siglen ty dflt = let rec sigrec_clausal_form siglen p_i = if Int.equal siglen 0 then (* is the default value typable with the expected type *) - let dflt_typ = unsafe_type_of env sigma (EConstr.of_constr dflt) in + let dflt_typ = unsafe_type_of env sigma dflt in try - let () = evdref := Evarconv.the_conv_x_leq env (EConstr.of_constr dflt_typ) (EConstr.of_constr p_i) !evdref in + let () = evdref := Evarconv.the_conv_x_leq env (EConstr.of_constr dflt_typ) p_i !evdref in let () = evdref := Evarconv.consider_remaining_unif_problems env !evdref in dflt with Evarconv.UnableToUnify _ -> error "Cannot solve a unification problem." else - let (a,p_i_minus_1) = match whd_beta_stack !evdref (EConstr.of_constr p_i) with - | (_sigS,[a;p]) -> (EConstr.Unsafe.to_constr a, EConstr.Unsafe.to_constr p) + let (a,p_i_minus_1) = match whd_beta_stack !evdref p_i with + | (_sigS,[a;p]) -> (a, p) | _ -> anomaly ~label:"sig_clausal_form" (Pp.str "should be a sigma type") in - let ev = Evarutil.e_new_evar env evdref (EConstr.of_constr a) in - let rty = beta_applist sigma (EConstr.of_constr p_i_minus_1,[ev]) in + let ev = Evarutil.e_new_evar env evdref a in + let rty = beta_applist sigma (p_i_minus_1,[ev]) in + let rty = EConstr.of_constr rty in let tuple_tail = sigrec_clausal_form (siglen-1) rty in let evopt = match EConstr.kind !evdref ev with Evar _ -> None | _ -> Some ev in match evopt with | Some w -> let w_type = unsafe_type_of env !evdref w in - if Evarconv.e_cumul env evdref (EConstr.of_constr w_type) (EConstr.of_constr a) then + let w_type = EConstr.of_constr w_type in + if Evarconv.e_cumul env evdref w_type a then let exist_term = Evarutil.evd_comb1 (Evd.fresh_global env) evdref sigdata.intro in - applist(exist_term,[a;p_i_minus_1;EConstr.Unsafe.to_constr w;tuple_tail]) + let exist_term = EConstr.of_constr exist_term in + applist(exist_term,[a;p_i_minus_1;w;tuple_tail]) else error "Cannot solve a unification problem." | None -> @@ -1218,7 +1230,7 @@ let sig_clausal_form env sigma sort_of_ty siglen ty dflt = error "Cannot solve a unification problem." in let scf = sigrec_clausal_form siglen ty in - !evdref, Evarutil.nf_evar !evdref scf + !evdref, EConstr.of_constr (Evarutil.nf_evar !evdref (EConstr.Unsafe.to_constr scf)) (* The problem is to build a destructor (a generalization of the predecessor) which, when applied to a term made of constructors @@ -1280,18 +1292,18 @@ let sig_clausal_form env sigma sort_of_ty siglen ty dflt = let make_iterated_tuple env sigma dflt (z,zty) = let (zty,rels) = minimal_free_rels_rec env sigma (z,zty) in - let sort_of_zty = get_sort_of env sigma (EConstr.of_constr zty) in + let sort_of_zty = get_sort_of env sigma zty in let sorted_rels = Int.Set.elements rels in let sigma, (tuple,tuplety) = List.fold_left (fun (sigma, t) -> make_tuple env sigma t) (sigma, (z,zty)) sorted_rels in - assert (closed0 tuplety); + assert (closed0 sigma tuplety); let n = List.length sorted_rels in let sigma, dfltval = sig_clausal_form env sigma sort_of_zty n tuplety dflt in sigma, (tuple,tuplety,dfltval) let rec build_injrec env sigma dflt c = function - | [] -> make_iterated_tuple env sigma dflt (c,unsafe_type_of env sigma (EConstr.of_constr c)) + | [] -> make_iterated_tuple env sigma dflt (c,EConstr.of_constr (unsafe_type_of env sigma c)) | ((sp,cnum),argnum)::l -> try let (cnum_nlams,cnum_env,kont) = descend_then env sigma c cnum in @@ -1326,40 +1338,44 @@ let inject_if_homogenous_dependent_pair ty = try let sigma = Tacmach.New.project gl in let eq,u,(t,t1,t2) = find_this_eq_data_decompose gl ty in - let t = EConstr.Unsafe.to_constr t in (* fetch the informations of the pair *) let ceq = Universes.constr_of_global Coqlib.glob_eq in + let ceq = EConstr.of_constr ceq in let sigTconstr () = (Coqlib.build_sigma_type()).Coqlib.typ in let existTconstr () = (Coqlib.build_sigma_type()).Coqlib.intro in (* check whether the equality deals with dep pairs or not *) - let eqTypeDest = fst (decompose_app t) in - if not (Globnames.is_global (sigTconstr()) eqTypeDest) then raise Exit; + let eqTypeDest = fst (decompose_app sigma t) in + if not (Termops.is_global sigma (sigTconstr()) eqTypeDest) then raise Exit; let hd1,ar1 = decompose_app_vect sigma t1 and hd2,ar2 = decompose_app_vect sigma t2 in - if not (Globnames.is_global (existTconstr()) hd1) then raise Exit; - if not (Globnames.is_global (existTconstr()) hd2) then raise Exit; - let ind,_ = try pf_apply find_mrectype gl (EConstr.of_constr ar1.(0)) with Not_found -> raise Exit in + let hd1 = EConstr.of_constr hd1 in + let hd2 = EConstr.of_constr hd2 in + let ar1 = Array.map EConstr.of_constr ar1 in + let ar2 = Array.map EConstr.of_constr ar2 in + if not (Termops.is_global sigma (existTconstr()) hd1) then raise Exit; + if not (Termops.is_global sigma (existTconstr()) hd2) then raise Exit; + let ind,_ = try pf_apply find_mrectype gl ar1.(0) with Not_found -> raise Exit in (* check if the user has declared the dec principle *) (* and compare the fst arguments of the dep pair *) (* Note: should work even if not an inductive type, but the table only *) (* knows inductive types *) if not (Ind_tables.check_scheme (!eq_dec_scheme_kind_name()) (fst ind) && - pf_apply is_conv gl (EConstr.of_constr ar1.(2)) (EConstr.of_constr ar2.(2))) then raise Exit; + pf_apply is_conv gl ar1.(2) ar2.(2)) then raise Exit; Coqlib.check_required_library ["Coq";"Logic";"Eqdep_dec"]; - let new_eq_args = [|pf_unsafe_type_of gl (EConstr.of_constr ar1.(3));ar1.(3);ar2.(3)|] in + let new_eq_args = [|EConstr.of_constr (pf_unsafe_type_of gl ar1.(3));ar1.(3);ar2.(3)|] in let inj2 = Coqlib.coq_constant "inj_pair2_eq_dec is missing" ["Logic";"Eqdep_dec"] "inj_pair2_eq_dec" in + let inj2 = EConstr.of_constr inj2 in let c, eff = find_scheme (!eq_dec_scheme_kind_name()) (Univ.out_punivs ind) in (* cut with the good equality and prove the requested goal *) tclTHENLIST [Proofview.tclEFFECTS eff; intro; onLastHyp (fun hyp -> - let hyp = EConstr.Unsafe.to_constr hyp in - tclTHENS (cut (EConstr.of_constr (mkApp (ceq,new_eq_args)))) - [clear [destVar hyp]; + tclTHENS (cut (mkApp (ceq,new_eq_args))) + [clear [destVar sigma hyp]; Proofview.V82.tactic (Tacmach.refine - (EConstr.of_constr (mkApp(inj2,[|ar1.(0);mkConst c;ar1.(1);ar1.(2);ar1.(3);ar2.(3);hyp|])))) + (mkApp(inj2,[|ar1.(0);mkConst c;ar1.(1);ar1.(2);ar1.(3);ar2.(3);hyp|]))) ])] with Exit -> Proofview.tclUNIT () @@ -1371,17 +1387,15 @@ let inject_if_homogenous_dependent_pair ty = let simplify_args env sigma t = (* Quick hack to reduce in arguments of eq only *) - match decompose_app t with - | eq, [t;c1;c2] -> applist (eq,[t;simpl env sigma (EConstr.of_constr c1);simpl env sigma (EConstr.of_constr c2)]) - | eq, [t1;c1;t2;c2] -> applist (eq,[t1;simpl env sigma (EConstr.of_constr c1);t2;simpl env sigma (EConstr.of_constr c2)]) + let simpl env sigma c = EConstr.of_constr (simpl env sigma c) in + match decompose_app sigma t with + | eq, [t;c1;c2] -> applist (eq,[t;simpl env sigma c1;simpl env sigma c2]) + | eq, [t1;c1;t2;c2] -> applist (eq,[t1;simpl env sigma c1;t2;simpl env sigma c2]) | _ -> t let inject_at_positions env sigma l2r (eq,_,(t,t1,t2)) eq_clause posns tac = - let t = EConstr.Unsafe.to_constr t in - let t1 = EConstr.Unsafe.to_constr t1 in - let t2 = EConstr.Unsafe.to_constr t2 in let e = next_ident_away eq_baseid (ids_of_context env) in - let e_env = push_named (LocalAssum (e,t)) env in + let e_env = push_named (nlocal_assum (e,t)) env in let evdref = ref sigma in let filter (cpath, t1', t2') = try @@ -1389,12 +1403,13 @@ let inject_at_positions env sigma l2r (eq,_,(t,t1,t2)) eq_clause posns tac = let sigma, (injbody,resty) = build_injector e_env !evdref t1' (mkVar e) cpath in let injfun = mkNamedLambda e t injbody in let sigma,congr = Evd.fresh_global env sigma eq.congr in + let congr = EConstr.of_constr congr in let pf = applist(congr,[t;resty;injfun;t1;t2]) in - let sigma, pf_typ = Typing.type_of env sigma (EConstr.of_constr pf) in + let sigma, pf_typ = Typing.type_of env sigma pf in + let pf_typ = EConstr.of_constr pf_typ in let inj_clause = apply_on_clause (pf,pf_typ) eq_clause in let pf = Clenvtac.clenv_value_cast_meta inj_clause in - let ty = simplify_args env sigma (EConstr.Unsafe.to_constr (clenv_type inj_clause)) in - let pf = EConstr.Unsafe.to_constr pf in + let ty = simplify_args env sigma (clenv_type inj_clause) in evdref := sigma; Some (pf, ty) with Failure _ -> None @@ -1406,9 +1421,9 @@ let inject_at_positions env sigma l2r (eq,_,(t,t1,t2)) eq_clause posns tac = Proofview.tclTHEN (Proofview.Unsafe.tclEVARS !evdref) (Tacticals.New.tclTHENFIRST (Proofview.tclIGNORE (Proofview.Monad.List.map - (fun (pf,ty) -> tclTHENS (cut (EConstr.of_constr ty)) - [inject_if_homogenous_dependent_pair (EConstr.of_constr ty); - Proofview.V82.tactic (Tacmach.refine (EConstr.of_constr pf))]) + (fun (pf,ty) -> tclTHENS (cut ty) + [inject_if_homogenous_dependent_pair ty; + Proofview.V82.tactic (Tacmach.refine pf)]) (if l2r then List.rev injectors else injectors))) (tac (List.length injectors))) @@ -1428,7 +1443,7 @@ let injEqThen tac l2r (eq,_,(t,t1,t2) as u) eq_clause = tclZEROMSG (str"Nothing to inject.") | Inr posns -> inject_at_positions env sigma l2r u eq_clause posns - (tac (EConstr.Unsafe.to_constr (clenv_value eq_clause))) + (tac (clenv_value eq_clause)) let get_previous_hyp_position id gl = let rec aux dest = function @@ -1451,10 +1466,10 @@ let injEq ?(old=false) with_evars clear_flag ipats = match ipats_style with | Some ipats -> Proofview.Goal.enter { enter = begin fun gl -> - let destopt = match kind_of_term c with + let sigma = project gl in + let destopt = match EConstr.kind sigma c with | Var id -> get_previous_hyp_position id gl | _ -> MoveLast in - let c = EConstr.of_constr c in let clear_tac = tclTRY (apply_clear_request clear_flag dft_clear_flag c) in (* Try should be removal if dependency were treated *) @@ -1488,10 +1503,10 @@ let decompEqThen ntac (lbeq,_,(t,t1,t2) as u) clause = | Inl (cpath, (_,dirn), _) -> discr_positions env sigma u clause cpath dirn | Inr [] -> (* Change: do not fail, simplify clear this trivial hyp *) - ntac (EConstr.Unsafe.to_constr (clenv_value clause)) 0 + ntac (clenv_value clause) 0 | Inr posns -> inject_at_positions env sigma true u clause posns - (ntac (EConstr.Unsafe.to_constr (clenv_value clause))) + (ntac (clenv_value clause)) end } let dEqThen with_evars ntac = function @@ -1500,7 +1515,6 @@ let dEqThen with_evars ntac = function let dEq with_evars = dEqThen with_evars (fun clear_flag c x -> - let c = EConstr.of_constr c in (apply_clear_request clear_flag (use_clear_hyp_by_default ()) c)) let intro_decomp_eq tac data (c, t) = @@ -1547,26 +1561,24 @@ let decomp_tuple_term env sigma c t = let rec decomprec inner_code ex exty = let iterated_decomp = try - let ex = EConstr.of_constr ex in let ({proj1=p1; proj2=p2}),(i,a,p,car,cdr) = find_sigma_data_decompose sigma ex in - let a = EConstr.Unsafe.to_constr a in - let p = EConstr.Unsafe.to_constr p in - let car = EConstr.Unsafe.to_constr car in - let cdr = EConstr.Unsafe.to_constr cdr in let car_code = applist (mkConstU (destConstRef p1,i),[a;p;inner_code]) and cdr_code = applist (mkConstU (destConstRef p2,i),[a;p;inner_code]) in - let cdrtyp = beta_applist sigma (EConstr.of_constr p,[EConstr.of_constr car]) in + let cdrtyp = beta_applist sigma (p,[car]) in + let cdrtyp = EConstr.of_constr cdrtyp in List.map (fun l -> ((car,a),car_code)::l) (decomprec cdr_code cdr cdrtyp) with Constr_matching.PatternMatchingFailure -> [] in [((ex,exty),inner_code)]::iterated_decomp in decomprec (mkRel 1) c t +let lambda_create env (a,b) = + mkLambda (named_hd env (EConstr.Unsafe.to_constr a) Anonymous, a, b) + let subst_tuple_term env sigma dep_pair1 dep_pair2 b = - let dep_pair1 = EConstr.Unsafe.to_constr dep_pair1 in - let dep_pair2 = EConstr.Unsafe.to_constr dep_pair2 in let sigma = Sigma.to_evar_map sigma in - let typ = get_type_of env sigma (EConstr.of_constr dep_pair1) in + let typ = get_type_of env sigma dep_pair1 in + let typ = EConstr.of_constr typ in (* We find all possible decompositions *) let decomps1 = decomp_tuple_term env sigma dep_pair1 typ in let decomps2 = decomp_tuple_term env sigma dep_pair2 typ in @@ -1581,15 +1593,18 @@ let subst_tuple_term env sigma dep_pair1 dep_pair2 b = (* We build the expected goal *) let abst_B = List.fold_right - (fun (e,t) body -> lambda_create env (t,subst_term sigma (EConstr.of_constr e) (EConstr.of_constr body))) e1_list b in - let pred_body = beta_applist sigma (EConstr.of_constr abst_B, List.map EConstr.of_constr proj_list) in + (fun (e,t) body -> lambda_create env (t,EConstr.of_constr (subst_term sigma e body))) e1_list b in + let pred_body = beta_applist sigma (abst_B,proj_list) in + let pred_body = EConstr.of_constr pred_body in let body = mkApp (lambda_create env (typ,pred_body),[|dep_pair1|]) in - let expected_goal = beta_applist sigma (EConstr.of_constr abst_B,List.map (fst %> EConstr.of_constr) e2_list) in + let expected_goal = beta_applist sigma (abst_B,List.map fst e2_list) in (* Simulate now the normalisation treatment made by Logic.mk_refgoals *) - let expected_goal = nf_betaiota sigma (EConstr.of_constr expected_goal) in + let expected_goal = EConstr.of_constr expected_goal in + let expected_goal = nf_betaiota sigma expected_goal in + let expected_goal = EConstr.of_constr expected_goal in (* Retype to get universes right *) - let sigma, expected_goal_ty = Typing.type_of env sigma (EConstr.of_constr expected_goal) in - let sigma, _ = Typing.type_of env sigma (EConstr.of_constr body) in + let sigma, expected_goal_ty = Typing.type_of env sigma expected_goal in + let sigma, _ = Typing.type_of env sigma body in Sigma.Unsafe.of_pair ((body, expected_goal), sigma) (* Like "replace" but decompose dependent equalities *) @@ -1598,16 +1613,14 @@ let subst_tuple_term env sigma dep_pair1 dep_pair2 b = (* on for further iterated sigma-tuples *) let cutSubstInConcl l2r eqn = - let eqn = EConstr.of_constr eqn in Proofview.Goal.nf_s_enter { s_enter = begin fun gl -> let env = Proofview.Goal.env gl in let sigma = Proofview.Goal.sigma gl in let (lbeq,u,(t,e1,e2)) = find_eq_data_decompose gl eqn in let typ = pf_concl gl in + let typ = EConstr.of_constr typ in let (e1,e2) = if l2r then (e1,e2) else (e2,e1) in let Sigma ((typ, expected), sigma, p) = subst_tuple_term env sigma e1 e2 typ in - let typ = EConstr.of_constr typ in - let expected = EConstr.of_constr expected in let tac = tclTHENFIRST (tclTHENLIST [ @@ -1620,16 +1633,14 @@ let cutSubstInConcl l2r eqn = end } let cutSubstInHyp l2r eqn id = - let eqn = EConstr.of_constr eqn in Proofview.Goal.nf_s_enter { s_enter = begin fun gl -> let env = Proofview.Goal.env gl in let sigma = Proofview.Goal.sigma gl in let (lbeq,u,(t,e1,e2)) = find_eq_data_decompose gl eqn in let typ = pf_get_hyp_typ id gl in + let typ = EConstr.of_constr typ in let (e1,e2) = if l2r then (e1,e2) else (e2,e1) in let Sigma ((typ, expected), sigma, p) = subst_tuple_term env sigma e1 e2 typ in - let typ = EConstr.of_constr typ in - let expected = EConstr.of_constr expected in let tac = tclTHENFIRST (tclTHENLIST [ @@ -1661,9 +1672,9 @@ let cutRewriteInHyp l2r eqn id = cutRewriteClause l2r eqn (Some id) let cutRewriteInConcl l2r eqn = cutRewriteClause l2r eqn None let substClause l2r c cls = - let c = EConstr.of_constr c in Proofview.Goal.enter { enter = begin fun gl -> let eq = pf_apply get_type_of gl c in + let eq = EConstr.of_constr eq in tclTHENS (cutSubstClause l2r eq cls) [Proofview.tclUNIT (); exact_no_check c] end } @@ -1707,7 +1718,7 @@ let restrict_to_eq_and_identity eq = (* compatibility *) not (is_global glob_identity eq) then raise Constr_matching.PatternMatchingFailure -exception FoundHyp of (Id.t * EConstr.constr * bool) +exception FoundHyp of (Id.t * constr * bool) (* tests whether hyp [c] is [x = t] or [t = x], [x] not occurring in [t] *) let is_eq_x gl x d = @@ -1779,7 +1790,7 @@ let subst_one_var dep_proof_ok x = user_err ~hdr:"Subst" (str "Cannot find any non-recursive equality over " ++ pr_id x ++ str".") - with FoundHyp (id, c, b) -> (id, EConstr.Unsafe.to_constr c, b) in + with FoundHyp res -> res in subst_one dep_proof_ok x res end } @@ -1811,15 +1822,14 @@ let subst_all ?(flags=default_subst_tactic_flags ()) () = let find_equations gl = let gl = Proofview.Goal.assume gl in let env = Proofview.Goal.env gl in + let sigma = project gl in let find_eq_data_decompose = find_eq_data_decompose gl in let select_equation_name decl = try let lbeq,u,(_,x,y) = find_eq_data_decompose (EConstr.of_constr (NamedDecl.get_type decl)) in - let x = EConstr.Unsafe.to_constr x in - let y = EConstr.Unsafe.to_constr y in let eq = Universes.constr_of_global_univ (lbeq.eq,u) in if flags.only_leibniz then restrict_to_eq_and_identity eq; - match kind_of_term x, kind_of_term y with + match EConstr.kind sigma x, EConstr.kind sigma y with | Var z, _ when not (is_evaluable env (EvalVarRef z)) -> Some (NamedDecl.get_id decl) | _, Var z when not (is_evaluable env (EvalVarRef z)) -> @@ -1842,14 +1852,12 @@ let subst_all ?(flags=default_subst_tactic_flags ()) () = let c = pf_get_hyp hyp gl |> NamedDecl.get_type in let c = EConstr.of_constr c in let _,_,(_,x,y) = find_eq_data_decompose c in - let x = EConstr.Unsafe.to_constr x in - let y = EConstr.Unsafe.to_constr y in (* J.F.: added to prevent failure on goal containing x=x as an hyp *) - if Term.eq_constr x y then Proofview.tclUNIT () else - match kind_of_term x, kind_of_term y with - | Var x', _ when not (occur_term sigma (EConstr.of_constr x) (EConstr.of_constr y)) && not (is_evaluable env (EvalVarRef x')) -> + if EConstr.eq_constr sigma x y then Proofview.tclUNIT () else + match EConstr.kind sigma x, EConstr.kind sigma y with + | Var x', _ when not (occur_term sigma x y) && not (is_evaluable env (EvalVarRef x')) -> subst_one flags.rewrite_dependent_proof x' (hyp,y,true) - | _, Var y' when not (occur_term sigma (EConstr.of_constr y) (EConstr.of_constr x)) && not (is_evaluable env (EvalVarRef y')) -> + | _, Var y' when not (occur_term sigma y x) && not (is_evaluable env (EvalVarRef y')) -> subst_one flags.rewrite_dependent_proof y' (hyp,x,false) | _ -> Proofview.tclUNIT () @@ -1866,19 +1874,18 @@ let subst_all ?(flags=default_subst_tactic_flags ()) () = or situations like "a = S b, b = S a", or also accidentally unfolding let-ins *) Proofview.Goal.nf_enter { enter = begin fun gl -> + let sigma = project gl in let find_eq_data_decompose = find_eq_data_decompose gl in let test (_,c) = try let c = EConstr.of_constr c in let lbeq,u,(_,x,y) = find_eq_data_decompose c in - let x = EConstr.Unsafe.to_constr x in - let y = EConstr.Unsafe.to_constr y in let eq = Universes.constr_of_global_univ (lbeq.eq,u) in if flags.only_leibniz then restrict_to_eq_and_identity eq; (* J.F.: added to prevent failure on goal containing x=x as an hyp *) - if Term.eq_constr x y then failwith "caught"; - match kind_of_term x with Var x -> x | _ -> - match kind_of_term y with Var y -> y | _ -> failwith "caught" + if EConstr.eq_constr sigma x y then failwith "caught"; + match EConstr.kind sigma x with Var x -> x | _ -> + match EConstr.kind sigma y with Var y -> y | _ -> failwith "caught" with Constr_matching.PatternMatchingFailure -> failwith "caught" in let test p = try Some (test p) with Failure _ -> None in let hyps = pf_hyps_types gl in @@ -1892,24 +1899,21 @@ let subst_all ?(flags=default_subst_tactic_flags ()) () = let cond_eq_term_left c t gl = try - let t = EConstr.of_constr t in let (_,x,_) = pi3 (find_eq_data_decompose gl t) in - if pf_conv_x gl (EConstr.of_constr c) x then true else failwith "not convertible" + if pf_conv_x gl c x then true else failwith "not convertible" with Constr_matching.PatternMatchingFailure -> failwith "not an equality" let cond_eq_term_right c t gl = try - let t = EConstr.of_constr t in let (_,_,x) = pi3 (find_eq_data_decompose gl t) in - if pf_conv_x gl (EConstr.of_constr c) x then false else failwith "not convertible" + if pf_conv_x gl c x then false else failwith "not convertible" with Constr_matching.PatternMatchingFailure -> failwith "not an equality" let cond_eq_term c t gl = try - let t = EConstr.of_constr t in let (_,x,y) = pi3 (find_eq_data_decompose gl t) in - if pf_conv_x gl (EConstr.of_constr c) x then true - else if pf_conv_x gl (EConstr.of_constr c) y then false + if pf_conv_x gl c x then true + else if pf_conv_x gl c y then false else failwith "not convertible" with Constr_matching.PatternMatchingFailure -> failwith "not an equality" @@ -1920,7 +1924,7 @@ let rewrite_assumption_cond cond_eq_term cl = let id = NamedDecl.get_id hyp in begin try - let dir = cond_eq_term (NamedDecl.get_type hyp) gl in + let dir = cond_eq_term (EConstr.of_constr (NamedDecl.get_type hyp)) gl in general_rewrite_clause dir false (mkVar id,NoBindings) cl with | Failure _ | UserError _ -> arec rest gl end @@ -1946,7 +1950,7 @@ let replace_term dir_opt c = (* Declare rewriting tactic for intro patterns "<-" and "->" *) let _ = - let gmr l2r with_evars tac c = general_rewrite_clause l2r with_evars (Miscops.map_with_bindings EConstr.Unsafe.to_constr tac) c in + let gmr l2r with_evars tac c = general_rewrite_clause l2r with_evars tac c in Hook.set Tactics.general_rewrite_clause gmr let _ = Hook.set Tactics.subst_one subst_one diff --git a/tactics/equality.mli b/tactics/equality.mli index 97f51ae20..5467b4af2 100644 --- a/tactics/equality.mli +++ b/tactics/equality.mli @@ -10,6 +10,7 @@ open Names open Term open Evd +open EConstr open Environ open Ind_tables open Locus @@ -60,30 +61,30 @@ val general_rewrite_clause : orientation -> evars_flag -> ?tac:(unit Proofview.tactic * conditions) -> constr with_bindings -> clause -> unit Proofview.tactic val general_multi_rewrite : - evars_flag -> (bool * multi * clear_flag * EConstr.constr with_bindings delayed_open) list -> + evars_flag -> (bool * multi * clear_flag * delayed_open_constr_with_bindings) list -> clause -> (unit Proofview.tactic * conditions) option -> unit Proofview.tactic val replace_in_clause_maybe_by : constr -> constr -> clause -> unit Proofview.tactic option -> unit Proofview.tactic val replace : constr -> constr -> unit Proofview.tactic val replace_by : constr -> constr -> unit Proofview.tactic -> unit Proofview.tactic -val discr : evars_flag -> EConstr.constr with_bindings -> unit Proofview.tactic +val discr : evars_flag -> constr with_bindings -> unit Proofview.tactic val discrConcl : unit Proofview.tactic val discrHyp : Id.t -> unit Proofview.tactic val discrEverywhere : evars_flag -> unit Proofview.tactic val discr_tac : evars_flag -> - EConstr.constr with_bindings destruction_arg option -> unit Proofview.tactic + constr with_bindings destruction_arg option -> unit Proofview.tactic val inj : intro_patterns option -> evars_flag -> - clear_flag -> EConstr.constr with_bindings -> unit Proofview.tactic + clear_flag -> constr with_bindings -> unit Proofview.tactic val injClause : intro_patterns option -> evars_flag -> - EConstr.constr with_bindings destruction_arg option -> unit Proofview.tactic + constr with_bindings destruction_arg option -> unit Proofview.tactic val injHyp : clear_flag -> Id.t -> unit Proofview.tactic val injConcl : unit Proofview.tactic val simpleInjClause : evars_flag -> - EConstr.constr with_bindings destruction_arg option -> unit Proofview.tactic + constr with_bindings destruction_arg option -> unit Proofview.tactic -val dEq : evars_flag -> EConstr.constr with_bindings destruction_arg option -> unit Proofview.tactic -val dEqThen : evars_flag -> (clear_flag -> constr -> int -> unit Proofview.tactic) -> EConstr.constr with_bindings destruction_arg option -> unit Proofview.tactic +val dEq : evars_flag -> constr with_bindings destruction_arg option -> unit Proofview.tactic +val dEqThen : evars_flag -> (clear_flag -> constr -> int -> unit Proofview.tactic) -> constr with_bindings destruction_arg option -> unit Proofview.tactic val make_iterated_tuple : env -> evar_map -> constr -> (constr * types) -> evar_map * (constr * constr * constr) @@ -96,8 +97,8 @@ val cutRewriteInConcl : bool -> constr -> unit Proofview.tactic val rewriteInHyp : bool -> constr -> Id.t -> unit Proofview.tactic val rewriteInConcl : bool -> constr -> unit Proofview.tactic -val discriminable : env -> evar_map -> EConstr.constr -> EConstr.constr -> bool -val injectable : env -> evar_map -> EConstr.constr -> EConstr.constr -> bool +val discriminable : env -> evar_map -> constr -> constr -> bool +val injectable : env -> evar_map -> constr -> constr -> bool (* Subst *) diff --git a/tactics/hints.ml b/tactics/hints.ml index 560e7e43d..c31e86383 100644 --- a/tactics/hints.ml +++ b/tactics/hints.ml @@ -737,7 +737,7 @@ let secvars_of_idset s = else p) s Id.Pred.empty let secvars_of_constr env c = - secvars_of_idset (global_vars_set env c) + secvars_of_idset (Environ.global_vars_set env c) let secvars_of_global env gr = secvars_of_idset (vars_of_global_reference env gr) diff --git a/tactics/inv.ml b/tactics/inv.ml index c66b356c7..ad2e2fa3b 100644 --- a/tactics/inv.ml +++ b/tactics/inv.ml @@ -113,7 +113,15 @@ let make_inv_predicate env evd indf realargs id status concl = if closed0 ti then (xi,ti,ai) else + let xi = EConstr.of_constr xi in + let ti = EConstr.of_constr ti in + let ai = EConstr.of_constr ai in let sigma, res = make_iterated_tuple env' !evd ai (xi,ti) in + let (xi, ti, ai) = res in + let xi = EConstr.Unsafe.to_constr xi in + let ti = EConstr.Unsafe.to_constr ti in + let ai = EConstr.Unsafe.to_constr ai in + let res = (xi, ti, ai) in evd := sigma; res in let eq_term = eqdata.Coqlib.eq in @@ -334,7 +342,7 @@ let remember_first_eq id x = if !x == MoveLast then x := MoveAfter id let projectAndApply as_mode thin avoid id eqname names depids = let subst_hyp l2r id = - tclTHEN (tclTRY(rewriteInConcl l2r (mkVar id))) + tclTHEN (tclTRY(rewriteInConcl l2r (EConstr.mkVar id))) (if thin then clear [id] else (remember_first_eq id eqname; tclIDTAC)) in let substHypIfVariable tac id = diff --git a/toplevel/auto_ind_decl.ml b/toplevel/auto_ind_decl.ml index 7759c400c..f8ca8343c 100644 --- a/toplevel/auto_ind_decl.ml +++ b/toplevel/auto_ind_decl.ml @@ -338,6 +338,8 @@ so from Ai we can find the the correct eq_Ai bl_ai or lb_ai *) (* used in the leib -> bool side*) let do_replace_lb mode lb_scheme_key aavoid narg p q = + let p = EConstr.of_constr p in + let q = EConstr.of_constr q in let avoid = Array.of_list aavoid in let do_arg v offset = try @@ -363,7 +365,7 @@ let do_replace_lb mode lb_scheme_key aavoid narg p q = ) in Proofview.Goal.nf_enter { enter = begin fun gl -> - let type_of_pq = Tacmach.New.of_old (fun gl -> pf_unsafe_type_of gl (EConstr.of_constr p)) gl in + let type_of_pq = Tacmach.New.of_old (fun gl -> pf_unsafe_type_of gl p) gl in let u,v = destruct_ind type_of_pq in let lb_type_of_p = try @@ -425,9 +427,11 @@ let do_replace_bl mode bl_scheme_key (ind,u as indu) aavoid narg lft rgt = let rec aux l1 l2 = match (l1,l2) with | (t1::q1,t2::q2) -> + let t1 = EConstr.of_constr t1 in + let t2 = EConstr.of_constr t2 in Proofview.Goal.enter { enter = begin fun gl -> - let tt1 = Tacmach.New.pf_unsafe_type_of gl (EConstr.of_constr t1) in - if Term.eq_constr t1 t2 then aux q1 q2 + let tt1 = Tacmach.New.pf_unsafe_type_of gl t1 in + if EConstr.eq_constr (Tacmach.New.project gl) t1 t2 then aux q1 q2 else ( let u,v = try destruct_ind tt1 (* trick so that the good sequence is returned*) @@ -923,7 +927,7 @@ let compute_dec_tact ind lnamesparrec nparrec = Equality.general_rewrite_bindings_in true Locus.AllOccurrences true false (List.hd !avoid) - ((mkVar (List.hd (List.tl !avoid))), + ((EConstr.mkVar (List.hd (List.tl !avoid))), NoBindings ) true; -- cgit v1.2.3 From 118ae18590dbc7d01cf34e0cd6133b1e34ef9090 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sat, 19 Nov 2016 00:20:30 +0100 Subject: Contradiction API using EConstr. --- engine/eConstr.ml | 1 + engine/eConstr.mli | 2 +- ltac/extratactics.ml4 | 2 +- plugins/omega/coq_omega.ml | 1 + tactics/contradiction.ml | 49 +++++++++++++++++++++++++--------------------- tactics/contradiction.mli | 3 ++- 6 files changed, 33 insertions(+), 25 deletions(-) diff --git a/engine/eConstr.ml b/engine/eConstr.ml index 1dd9d0c00..629be8e4b 100644 --- a/engine/eConstr.ml +++ b/engine/eConstr.ml @@ -99,6 +99,7 @@ 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) diff --git a/engine/eConstr.mli b/engine/eConstr.mli index e6270fa78..cb671154c 100644 --- a/engine/eConstr.mli +++ b/engine/eConstr.mli @@ -69,7 +69,7 @@ val mkInd : inductive -> t val mkIndU : pinductive -> t val mkConstruct : constructor -> t val mkConstructU : pconstructor -> t -(* val mkConstructUi : pinductive * int -> t *) +val mkConstructUi : pinductive * int -> t val mkCase : case_info * t * t * t array -> t val mkFix : (t, t) pfixpoint -> t val mkCoFix : (t, t) pcofixpoint -> t diff --git a/ltac/extratactics.ml4 b/ltac/extratactics.ml4 index bcfa13c79..f87aa407d 100644 --- a/ltac/extratactics.ml4 +++ b/ltac/extratactics.ml4 @@ -185,7 +185,7 @@ END open Contradiction TACTIC EXTEND absurd - [ "absurd" constr(c) ] -> [ absurd c ] + [ "absurd" constr(c) ] -> [ absurd (EConstr.of_constr c) ] END let onSomeWithHoles tac = function diff --git a/plugins/omega/coq_omega.ml b/plugins/omega/coq_omega.ml index 35d763ccc..f39549514 100644 --- a/plugins/omega/coq_omega.ml +++ b/plugins/omega/coq_omega.ml @@ -1092,6 +1092,7 @@ let replay_history tactic_normalisation = Lazy.force coq_Gt; Lazy.force coq_Gt |]) in + let not_sup_sup = EConstr.of_constr not_sup_sup in Tacticals.New.tclTHENS (Tacticals.New.tclTHENLIST [ unfold sp_Zle; diff --git a/tactics/contradiction.ml b/tactics/contradiction.ml index a92b14dbe..596f1a759 100644 --- a/tactics/contradiction.ml +++ b/tactics/contradiction.ml @@ -7,6 +7,7 @@ (************************************************************************) open Term +open EConstr open Hipattern open Tactics open Coqlib @@ -19,6 +20,7 @@ module NamedDecl = Context.Named.Declaration (* Absurd *) let mk_absurd_proof t = + let build_coq_not () = EConstr.of_constr (build_coq_not ()) in let id = Namegen.default_dependent_ident in mkLambda (Names.Name id,mkApp(build_coq_not (),[|t|]), mkLambda (Names.Name id,t,mkApp (mkRel 2,[|mkRel 1|]))) @@ -28,13 +30,13 @@ let absurd c = let sigma = Proofview.Goal.sigma gl in let env = Proofview.Goal.env gl in let sigma = Sigma.to_evar_map sigma in - let j = Retyping.get_judgment_of env sigma (EConstr.of_constr c) in + let j = Retyping.get_judgment_of env sigma c in let sigma, j = Coercion.inh_coerce_to_sort Loc.ghost env sigma j in - let t = EConstr.Unsafe.to_constr j.Environ.utj_val in + let t = j.Environ.utj_val in let tac = Tacticals.New.tclTHENLIST [ elim_type (EConstr.of_constr (build_coq_False ())); - Simple.apply (EConstr.of_constr (mk_absurd_proof t)) + Simple.apply (mk_absurd_proof t) ] in Sigma.Unsafe.of_pair (tac, sigma) end } @@ -49,7 +51,7 @@ let use_negated_unit_or_eq_type () = Flags.version_strictly_greater Flags.V8_5 let filter_hyp f tac = let rec seek = function | [] -> Proofview.tclZERO Not_found - | d::rest when f (NamedDecl.get_type d) -> tac (NamedDecl.get_id d) + | d::rest when f (EConstr.of_constr (NamedDecl.get_type d)) -> tac (NamedDecl.get_id d) | _::rest -> seek rest in Proofview.Goal.enter { enter = begin fun gl -> let hyps = Proofview.Goal.hyps (Proofview.Goal.assume gl) in @@ -65,31 +67,33 @@ let contradiction_context = | d :: rest -> let id = NamedDecl.get_id d in let typ = nf_evar sigma (NamedDecl.get_type d) in - let typ = whd_all env sigma (EConstr.of_constr typ) in - if is_empty_type sigma (EConstr.of_constr typ) then - simplest_elim (EConstr.mkVar id) - else match kind_of_term typ with - | Prod (na,t,u) when is_empty_type sigma (EConstr.of_constr u) -> + let typ = EConstr.of_constr typ in + let typ = whd_all env sigma typ in + let typ = EConstr.of_constr typ in + if is_empty_type sigma typ then + simplest_elim (mkVar id) + else match EConstr.kind sigma typ with + | Prod (na,t,u) when is_empty_type sigma u -> let is_unit_or_eq = - if use_negated_unit_or_eq_type () then match_with_unit_or_eq_type sigma (EConstr.of_constr t) + if use_negated_unit_or_eq_type () then match_with_unit_or_eq_type sigma t else None in Tacticals.New.tclORELSE (match is_unit_or_eq with | Some _ -> - let hd,args = decompose_app t in - let (ind,_ as indu) = destInd hd in + let hd,args = decompose_app sigma t in + let (ind,_ as indu) = destInd sigma hd in let nparams = Inductiveops.inductive_nparams_env env ind in let params = Util.List.firstn nparams args in let p = applist ((mkConstructUi (indu,1)), params) in (* Checking on the fly that it type-checks *) - simplest_elim (EConstr.mkApp (EConstr.mkVar id,[|EConstr.of_constr p|])) + simplest_elim (mkApp (mkVar id,[|p|])) | None -> Tacticals.New.tclZEROMSG (Pp.str"Not a negated unit type.")) (Proofview.tclORELSE (Proofview.Goal.enter { enter = begin fun gl -> let is_conv_leq = Tacmach.New.pf_apply is_conv_leq gl in - filter_hyp (fun typ -> is_conv_leq (EConstr.of_constr typ) (EConstr.of_constr t)) - (fun id' -> simplest_elim (EConstr.mkApp (EConstr.mkVar id,[|EConstr.mkVar id'|]))) + filter_hyp (fun typ -> is_conv_leq typ t) + (fun id' -> simplest_elim (mkApp (mkVar id,[|mkVar id'|]))) end }) begin function (e, info) -> match e with | Not_found -> seek_neg rest @@ -102,10 +106,9 @@ let contradiction_context = end } let is_negation_of env sigma typ t = - match kind_of_term (whd_all env sigma t) with + match EConstr.kind sigma (EConstr.of_constr (whd_all env sigma t)) with | Prod (na,t,u) -> - let u = nf_evar sigma u in - is_empty_type sigma (EConstr.of_constr u) && is_conv_leq env sigma (EConstr.of_constr typ) (EConstr.of_constr t) + is_empty_type sigma u && is_conv_leq env sigma typ t | _ -> false let contradiction_term (c,lbind as cl) = @@ -114,8 +117,10 @@ let contradiction_term (c,lbind as cl) = let env = Proofview.Goal.env gl in let type_of = Tacmach.New.pf_unsafe_type_of gl in let typ = type_of c in - let _, ccl = splay_prod env sigma (EConstr.of_constr typ) in - if is_empty_type sigma (EConstr.of_constr ccl) then + let typ = EConstr.of_constr typ in + let _, ccl = splay_prod env sigma typ in + let ccl = EConstr.of_constr ccl in + if is_empty_type sigma ccl then Tacticals.New.tclTHEN (elim false None cl None) (Tacticals.New.tclTRY assumption) @@ -123,8 +128,8 @@ let contradiction_term (c,lbind as cl) = Proofview.tclORELSE begin if lbind = NoBindings then - filter_hyp (fun c -> is_negation_of env sigma typ (EConstr.of_constr c)) - (fun id -> simplest_elim (EConstr.mkApp (EConstr.mkVar id,[|c|]))) + filter_hyp (fun c -> is_negation_of env sigma typ c) + (fun id -> simplest_elim (mkApp (mkVar id,[|c|]))) else Proofview.tclZERO Not_found end diff --git a/tactics/contradiction.mli b/tactics/contradiction.mli index 5cc4b2e01..510b135b0 100644 --- a/tactics/contradiction.mli +++ b/tactics/contradiction.mli @@ -7,7 +7,8 @@ (************************************************************************) open Term +open EConstr open Misctypes val absurd : constr -> unit Proofview.tactic -val contradiction : EConstr.constr with_bindings option -> unit Proofview.tactic +val contradiction : constr with_bindings option -> unit Proofview.tactic -- cgit v1.2.3 From db252cb87e9c63f400fd4fddd2d771df3160d592 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sat, 19 Nov 2016 01:07:35 +0100 Subject: Inv API using EConstr. --- engine/termops.ml | 4 +-- engine/termops.mli | 4 +-- ltac/rewrite.ml | 7 ++++- plugins/extraction/extraction.ml | 1 + pretyping/nativenorm.ml | 2 ++ pretyping/recordops.ml | 4 ++- pretyping/reductionops.ml | 11 ++++---- pretyping/reductionops.mli | 6 ++--- pretyping/vnorm.ml | 2 ++ tactics/contradiction.ml | 1 - tactics/hipattern.ml | 3 ++- tactics/hipattern.mli | 2 +- tactics/inv.ml | 55 +++++++++++++++++++--------------------- tactics/inv.mli | 3 ++- tactics/tactics.ml | 6 ++--- 15 files changed, 61 insertions(+), 50 deletions(-) diff --git a/engine/termops.ml b/engine/termops.ml index c2d862f00..4ab9f0733 100644 --- a/engine/termops.ml +++ b/engine/termops.ml @@ -224,8 +224,8 @@ let mkProd_wo_LetIn decl c = | LocalAssum (na,t) -> mkProd (na, EConstr.of_constr t, c) | LocalDef (_,b,_) -> Vars.subst1 (EConstr.of_constr 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 diff --git a/engine/termops.mli b/engine/termops.mli index 013efcbcb..1c14e6c89 100644 --- a/engine/termops.mli +++ b/engine/termops.mli @@ -46,8 +46,8 @@ val rel_list : int -> int -> EConstr.constr list (** iterators/destructors on terms *) val mkProd_or_LetIn : Context.Rel.Declaration.t -> types -> types val mkProd_wo_LetIn : Context.Rel.Declaration.t -> EConstr.types -> EConstr.types -val it_mkProd : types -> (Name.t * types) list -> types -val it_mkLambda : constr -> (Name.t * types) list -> constr +val it_mkProd : EConstr.types -> (Name.t * EConstr.types) list -> EConstr.types +val it_mkLambda : EConstr.constr -> (Name.t * EConstr.types) list -> EConstr.constr val it_mkProd_or_LetIn : types -> Context.Rel.t -> types val it_mkProd_wo_LetIn : EConstr.types -> Context.Rel.t -> EConstr.types val it_mkLambda_or_LetIn : constr -> Context.Rel.t -> constr diff --git a/ltac/rewrite.ml b/ltac/rewrite.ml index 0d279ae93..f2ffb0413 100644 --- a/ltac/rewrite.ml +++ b/ltac/rewrite.ml @@ -39,6 +39,10 @@ open Context.Named.Declaration module NamedDecl = Context.Named.Declaration module RelDecl = Context.Rel.Declaration +let local_assum (na, t) = + let inj = EConstr.Unsafe.to_constr in + RelDecl.LocalAssum (na, inj t) + (** Typeclass-based generalized rewriting. *) (** Constants used by the tactic. *) @@ -531,7 +535,8 @@ let decompose_applied_relation env sigma (c,l) = | Some c -> c | None -> let ctx,t' = Reductionops.splay_prod env sigma (EConstr.of_constr ctype) in (* Search for underlying eq *) - match find_rel (it_mkProd_or_LetIn t' (List.map (fun (n,t) -> LocalAssum (n, t)) ctx)) with + let t' = EConstr.Unsafe.to_constr t' in + match find_rel (it_mkProd_or_LetIn t' (List.map (fun (n,t) -> local_assum (n, t)) ctx)) with | Some c -> c | None -> error "Cannot find an homogeneous relation to rewrite." diff --git a/plugins/extraction/extraction.ml b/plugins/extraction/extraction.ml index 6559aeb08..e9fd40722 100644 --- a/plugins/extraction/extraction.ml +++ b/plugins/extraction/extraction.ml @@ -325,6 +325,7 @@ and extract_type_scheme env db c p = extract_type_scheme (push_rel_assum (n,t) env) db d (p-1) | _ -> let rels = fst (splay_prod env none (EConstr.of_constr (type_of env c))) in + let rels = List.map (on_snd EConstr.Unsafe.to_constr) rels in let env = push_rels_assum rels env in let eta_args = List.rev_map mkRel (List.interval 1 p) in extract_type env db 0 (lift p c) eta_args diff --git a/pretyping/nativenorm.ml b/pretyping/nativenorm.ml index ff3424c44..cdaa4e9ee 100644 --- a/pretyping/nativenorm.ml +++ b/pretyping/nativenorm.ml @@ -99,6 +99,8 @@ let build_branches_type env sigma (mind,_ as _ind) mib mip u params dep p = let build_one_branch i cty = let typi = type_constructor mind mib u cty params in let decl,indapp = Reductionops.splay_prod env sigma (EConstr.of_constr typi) in + let decl = List.map (on_snd EConstr.Unsafe.to_constr) decl in + let indapp = EConstr.Unsafe.to_constr indapp in let decl_with_letin,_ = decompose_prod_assum typi in let ind,cargs = find_rectype_a env indapp in let nparams = Array.length params in diff --git a/pretyping/recordops.ml b/pretyping/recordops.ml index f09f3221d..3230f92da 100644 --- a/pretyping/recordops.ml +++ b/pretyping/recordops.ml @@ -203,7 +203,8 @@ let compute_canonical_projections warn (con,ind) = let ctx = Univ.ContextSet.of_context ctx in let c = Environ.constant_value_in env (con,u) in let lt,t = Reductionops.splay_lam env Evd.empty (EConstr.of_constr c) in - let lt = List.rev_map snd lt in + let t = EConstr.Unsafe.to_constr t in + let lt = List.rev_map (snd %> EConstr.Unsafe.to_constr) lt in let args = snd (decompose_app t) in let { s_EXPECTEDPARAM = p; s_PROJ = lpj; s_PROJKIND = kl } = lookup_structure ind in @@ -303,6 +304,7 @@ let check_and_decompose_canonical_structure ref = | Some vc -> vc | None -> error_not_structure ref in let body = snd (splay_lam (Global.env()) Evd.empty (EConstr.of_constr vc)) (** FIXME *) in + let body = EConstr.Unsafe.to_constr body in let f,args = match kind_of_term body with | App (f,args) -> f,args | _ -> error_not_structure ref in diff --git a/pretyping/reductionops.ml b/pretyping/reductionops.ml index 0dd615bfb..480ec2319 100644 --- a/pretyping/reductionops.ml +++ b/pretyping/reductionops.ml @@ -1453,13 +1453,13 @@ let hnf_lam_applist env sigma t nl = List.fold_left (fun acc t -> hnf_lam_app env sigma (EConstr.of_constr acc) t) (EConstr.Unsafe.to_constr t) nl let bind_assum (na, t) = - let inj = EConstr.Unsafe.to_constr in - (na, inj t) + (na, t) let splay_prod env sigma = let rec decrec env m c = let t = whd_all env sigma c in - match EConstr.kind sigma (EConstr.of_constr t) with + let t = EConstr.of_constr t in + match EConstr.kind sigma t with | Prod (n,a,c0) -> decrec (push_rel (local_assum (n,a)) env) (bind_assum (n,a)::m) c0 @@ -1470,7 +1470,8 @@ let splay_prod env sigma = let splay_lam env sigma = let rec decrec env m c = let t = whd_all env sigma c in - match EConstr.kind sigma (EConstr.of_constr t) with + let t = EConstr.of_constr t in + match EConstr.kind sigma t with | Lambda (n,a,c0) -> decrec (push_rel (local_assum (n,a)) env) (bind_assum (n,a)::m) c0 @@ -1498,7 +1499,7 @@ let splay_prod_assum env sigma = let splay_arity env sigma c = let l, c = splay_prod env sigma c in - match EConstr.kind sigma (EConstr.of_constr c) with + match EConstr.kind sigma c with | Sort s -> l,s | _ -> invalid_arg "splay_arity" diff --git a/pretyping/reductionops.mli b/pretyping/reductionops.mli index 864b1a625..e67fad3fd 100644 --- a/pretyping/reductionops.mli +++ b/pretyping/reductionops.mli @@ -213,9 +213,9 @@ val hnf_lam_app : env -> evar_map -> EConstr.t -> EConstr.t -> constr val hnf_lam_appvect : env -> evar_map -> EConstr.t -> EConstr.t array -> constr val hnf_lam_applist : env -> evar_map -> EConstr.t -> EConstr.t list -> constr -val splay_prod : env -> evar_map -> EConstr.t -> (Name.t * constr) list * constr -val splay_lam : env -> evar_map -> EConstr.t -> (Name.t * constr) list * constr -val splay_arity : env -> evar_map -> EConstr.t -> (Name.t * constr) list * sorts +val splay_prod : env -> evar_map -> EConstr.t -> (Name.t * EConstr.constr) list * EConstr.constr +val splay_lam : env -> evar_map -> EConstr.t -> (Name.t * EConstr.constr) list * EConstr.constr +val splay_arity : env -> evar_map -> EConstr.t -> (Name.t * EConstr.constr) list * sorts val sort_of_arity : env -> evar_map -> EConstr.t -> sorts val splay_prod_n : env -> evar_map -> int -> EConstr.t -> Context.Rel.t * constr val splay_lam_n : env -> evar_map -> int -> EConstr.t -> Context.Rel.t * constr diff --git a/pretyping/vnorm.ml b/pretyping/vnorm.ml index 60f99fd3d..31693d82f 100644 --- a/pretyping/vnorm.ml +++ b/pretyping/vnorm.ml @@ -106,6 +106,8 @@ let build_branches_type env sigma (mind,_ as _ind) mib mip u params dep p = let build_one_branch i cty = let typi = type_constructor mind mib u cty params in let decl,indapp = Reductionops.splay_prod env sigma (EConstr.of_constr typi) in + let decl = List.map (on_snd EConstr.Unsafe.to_constr) decl in + let indapp = EConstr.Unsafe.to_constr indapp in let decl_with_letin,_ = decompose_prod_assum typi in let ((ind,u),cargs) = find_rectype_a env indapp in let nparams = Array.length params in diff --git a/tactics/contradiction.ml b/tactics/contradiction.ml index 596f1a759..2d5c28eba 100644 --- a/tactics/contradiction.ml +++ b/tactics/contradiction.ml @@ -119,7 +119,6 @@ let contradiction_term (c,lbind as cl) = let typ = type_of c in let typ = EConstr.of_constr typ in let _, ccl = splay_prod env sigma typ in - let ccl = EConstr.of_constr ccl in if is_empty_type sigma ccl then Tacticals.New.tclTHEN (elim false None cl None) diff --git a/tactics/hipattern.ml b/tactics/hipattern.ml index 6681e5e49..36ed589b9 100644 --- a/tactics/hipattern.ml +++ b/tactics/hipattern.ml @@ -470,10 +470,11 @@ let match_eq_nf gls eqn (ref, hetero) = let n = if hetero then 4 else 3 in let args = List.init n (fun i -> mkGPatVar ("X" ^ string_of_int (i + 1))) in let pat = mkPattern (mkGAppRef ref args) in + let pf_whd_all gls c = EConstr.of_constr (pf_whd_all gls (EConstr.of_constr c)) in match Id.Map.bindings (pf_matches gls pat eqn) with | [(m1,t);(m2,x);(m3,y)] -> assert (Id.equal m1 meta1 && Id.equal m2 meta2 && Id.equal m3 meta3); - (t,pf_whd_all gls (EConstr.of_constr x),pf_whd_all gls (EConstr.of_constr y)) + (EConstr.of_constr t,pf_whd_all gls x,pf_whd_all gls y) | _ -> anomaly ~label:"match_eq" (Pp.str "an eq pattern should match 3 terms") let dest_nf_eq gls eqn = diff --git a/tactics/hipattern.mli b/tactics/hipattern.mli index 094d62df6..c061c50f0 100644 --- a/tactics/hipattern.mli +++ b/tactics/hipattern.mli @@ -146,7 +146,7 @@ val is_matching_sigma : evar_map -> constr -> bool val match_eqdec : evar_map -> constr -> bool * Constr.constr * Constr.constr * Constr.constr * Constr.constr (** Match an equality up to conversion; returns [(eq,t1,t2)] in normal form *) -val dest_nf_eq : ([ `NF ], 'r) Proofview.Goal.t -> constr -> (Constr.constr * Constr.constr * Constr.constr) +val dest_nf_eq : ([ `NF ], 'r) Proofview.Goal.t -> constr -> (constr * constr * constr) (** Match a negation *) val is_matching_not : evar_map -> constr -> bool diff --git a/tactics/inv.ml b/tactics/inv.ml index ad2e2fa3b..37c82ff64 100644 --- a/tactics/inv.ml +++ b/tactics/inv.ml @@ -12,8 +12,9 @@ open Util open Names open Nameops open Term -open Vars open Termops +open EConstr +open Vars open Namegen open Environ open Inductiveops @@ -62,10 +63,10 @@ let var_occurs_in_pf gl id = *) -type inversion_status = Dep of constr option | NoDep +type inversion_status = Dep of EConstr.constr option | NoDep let compute_eqn env sigma n i ai = - (mkRel (n-i),get_type_of env sigma (EConstr.of_constr (mkRel (n-i)))) + (mkRel (n-i),EConstr.of_constr (get_type_of env sigma (mkRel (n-i)))) let make_inv_predicate env evd indf realargs id status concl = let nrealargs = List.length realargs in @@ -76,7 +77,7 @@ let make_inv_predicate env evd indf realargs id status concl = let hyps_arity,_ = get_arity env indf in (hyps_arity,concl) | Dep dflt_concl -> - if not (occur_var env !evd id (EConstr.of_constr concl)) then + if not (occur_var env !evd id concl) then user_err ~hdr:"make_inv_predicate" (str "Current goal does not depend on " ++ pr_id id ++ str"."); (* We abstract the conclusion of goal with respect to @@ -86,13 +87,14 @@ let make_inv_predicate env evd indf realargs id status concl = match dflt_concl with | Some concl -> concl (*assumed it's some [x1..xn,H:I(x1..xn)]C*) | None -> - let sort = get_sort_family_of env !evd (EConstr.of_constr concl) in + let sort = get_sort_family_of env !evd concl in let sort = Evarutil.evd_comb1 (Evd.fresh_sort_in_family env) evd sort in let p = make_arity env true indf sort in + let p = EConstr.of_constr p in let evd',(p,ptyp) = Unification.abstract_list_all env - !evd (EConstr.of_constr p) (EConstr.of_constr concl) (List.map EConstr.of_constr realargs@[EConstr.mkVar id]) - in evd := evd'; EConstr.Unsafe.to_constr p in - let hyps,bodypred = decompose_lam_n_assum (nrealargs+1) pred in + !evd p concl (realargs@[mkVar id]) + in evd := evd'; p in + let hyps,bodypred = decompose_lam_n_assum !evd (nrealargs+1) pred in (* We lift to make room for the equations *) (hyps,lift nrealargs bodypred) in @@ -110,34 +112,28 @@ let make_inv_predicate env evd indf realargs id status concl = let ai = lift nhyps ai in let (xi, ti) = compute_eqn env' !evd nhyps n ai in let (lhs,eqnty,rhs) = - if closed0 ti then + if closed0 !evd ti then (xi,ti,ai) else - let xi = EConstr.of_constr xi in - let ti = EConstr.of_constr ti in - let ai = EConstr.of_constr ai in let sigma, res = make_iterated_tuple env' !evd ai (xi,ti) in - let (xi, ti, ai) = res in - let xi = EConstr.Unsafe.to_constr xi in - let ti = EConstr.Unsafe.to_constr ti in - let ai = EConstr.Unsafe.to_constr ai in - let res = (xi, ti, ai) in evd := sigma; res in let eq_term = eqdata.Coqlib.eq in let eq = Evarutil.evd_comb1 (Evd.fresh_global env) evd eq_term in + let eq = EConstr.of_constr eq in let eqn = applist (eq,[eqnty;lhs;rhs]) in let eqns = (Anonymous, lift n eqn) :: eqns in let refl_term = eqdata.Coqlib.refl in let refl_term = Evarutil.evd_comb1 (Evd.fresh_global env) evd refl_term in + let refl_term = EConstr.of_constr refl_term in let refl = mkApp (refl_term, [|eqnty; rhs|]) in - let _ = Evarutil.evd_comb1 (Typing.type_of env) evd (EConstr.of_constr refl) in + let _ = Evarutil.evd_comb1 (Typing.type_of env) evd refl in let args = refl :: args in build_concl eqns args (succ n) restlist in let (newconcl, args) = build_concl [] [] 0 realargs in - let predicate = it_mkLambda_or_LetIn_name env newconcl hyps in - let _ = Evarutil.evd_comb1 (Typing.type_of env) evd (EConstr.of_constr predicate) in + let predicate = it_mkLambda_or_LetIn newconcl (name_context env hyps) in + let _ = Evarutil.evd_comb1 (Typing.type_of env) evd predicate in (* OK - this predicate should now be usable by res_elimination_then to do elimination on the conclusion. *) predicate, args @@ -347,10 +343,11 @@ let projectAndApply as_mode thin avoid id eqname names depids = in let substHypIfVariable tac id = Proofview.Goal.nf_enter { enter = begin fun gl -> + let sigma = project gl in (** We only look at the type of hypothesis "id" *) let hyp = pf_nf_evar gl (pf_get_hyp_typ id (Proofview.Goal.assume gl)) in let (t,t1,t2) = Hipattern.dest_nf_eq gl (EConstr.of_constr hyp) in - match (kind_of_term t1, kind_of_term t2) with + match (EConstr.kind sigma t1, EConstr.kind sigma t2) with | Var id1, _ -> generalizeRewriteIntros as_mode (subst_hyp true id) depids id1 | _, Var id2 -> generalizeRewriteIntros as_mode (subst_hyp false id) depids id2 | _ -> tac id @@ -444,42 +441,42 @@ let raw_inversion inv_kind id status names = let sigma = Sigma.to_evar_map sigma in let env = Proofview.Goal.env gl in let concl = Proofview.Goal.concl gl in + let concl = EConstr.of_constr concl in let c = mkVar id in let (ind, t) = - try pf_apply Tacred.reduce_to_atomic_ind gl (EConstr.of_constr (pf_unsafe_type_of gl (EConstr.of_constr c))) + try pf_apply Tacred.reduce_to_atomic_ind gl (EConstr.of_constr (pf_unsafe_type_of gl c)) with UserError _ -> let msg = str "The type of " ++ pr_id id ++ str " is not inductive." in CErrors.user_err msg in let IndType (indf,realargs) = find_rectype env sigma t in + let realargs = List.map EConstr.of_constr realargs in let evdref = ref sigma in let (elim_predicate, args) = make_inv_predicate env evdref indf realargs id status concl in let sigma = !evdref in let (cut_concl,case_tac) = - if status != NoDep && (dependent sigma (EConstr.of_constr c) (EConstr.of_constr concl)) then - Reduction.beta_appvect elim_predicate (Array.of_list (realargs@[c])), + if status != NoDep && (dependent sigma c concl) then + Reductionops.beta_applist sigma (elim_predicate, realargs@[c]), case_then_using else - Reduction.beta_appvect elim_predicate (Array.of_list realargs), + Reductionops.beta_applist sigma (elim_predicate, realargs), case_nodep_then_using in let cut_concl = EConstr.of_constr cut_concl in let refined id = let prf = mkApp (mkVar id, args) in - let prf = EConstr.of_constr prf in Refine.refine { run = fun h -> Sigma (prf, h, Sigma.refl) } in let neqns = List.length realargs in let as_mode = names != None in - let elim_predicate = EConstr.of_constr elim_predicate in let tac = (tclTHENS (assert_before Anonymous cut_concl) [case_tac names (introCaseAssumsThen false (* ApplyOn not supported by inversion *) (rewrite_equations_tac as_mode inv_kind id neqns)) - (Some elim_predicate) ind (EConstr.of_constr c,t); + (Some elim_predicate) ind (c,t); onLastHypId (fun id -> tclTHEN (refined id) reflexivity)]) in Sigma.Unsafe.of_pair (tac, sigma) @@ -513,7 +510,7 @@ let inv k = inv_gen k NoDep let inv_tac id = inv FullInversion None (NamedHyp id) let inv_clear_tac id = inv FullInversionClear None (NamedHyp id) -let dinv k c = inv_gen k (Dep (Option.map EConstr.Unsafe.to_constr c)) +let dinv k c = inv_gen k (Dep c) let dinv_tac id = dinv FullInversion None None (NamedHyp id) let dinv_clear_tac id = dinv FullInversionClear None None (NamedHyp id) diff --git a/tactics/inv.mli b/tactics/inv.mli index 6bb2b7282..446a62f6d 100644 --- a/tactics/inv.mli +++ b/tactics/inv.mli @@ -8,6 +8,7 @@ open Names open Term +open EConstr open Misctypes open Tactypes @@ -20,7 +21,7 @@ val inv_clause : val inv : inversion_kind -> or_and_intro_pattern option -> quantified_hypothesis -> unit Proofview.tactic -val dinv : inversion_kind -> EConstr.constr option -> +val dinv : inversion_kind -> constr option -> or_and_intro_pattern option -> quantified_hypothesis -> unit Proofview.tactic val inv_tac : Id.t -> unit Proofview.tactic diff --git a/tactics/tactics.ml b/tactics/tactics.ml index b9da11021..bae1ad48c 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -4367,6 +4367,7 @@ let clear_unselected_context id inhyps cls = let use_bindings env sigma elim must_be_closed (c,lbind) typ = let sigma = Sigma.to_evar_map sigma in + let typ = EConstr.of_constr typ in let typ = if elim == None then (* w/o an scheme, the term has to be applied at least until @@ -4374,7 +4375,7 @@ let use_bindings env sigma elim must_be_closed (c,lbind) typ = known only by pattern-matching, as in the case of a term of the form "nat_rect ?A ?o ?s n", with ?A to be inferred by matching. *) - let sign,t = splay_prod env sigma (EConstr.of_constr typ) in it_mkProd t sign + let sign,t = splay_prod env sigma typ in it_mkProd t sign else (* Otherwise, we exclude the case of an induction argument in an explicitly functional type. Henceforth, we can complete the @@ -4384,7 +4385,6 @@ let use_bindings env sigma elim must_be_closed (c,lbind) typ = typ in let rec find_clause typ = try - let typ = EConstr.of_constr typ in let indclause = make_clenv_binding env sigma (c,typ) lbind in if must_be_closed && occur_meta indclause.evd (clenv_value indclause) then error "Need a fully applied argument."; @@ -4392,7 +4392,7 @@ let use_bindings env sigma elim must_be_closed (c,lbind) typ = let (sigma, c) = pose_all_metas_as_evars env indclause.evd (clenv_value indclause) in Sigma.Unsafe.of_pair (c, sigma) with e when catchable_exception e -> - try find_clause (try_red_product env sigma (EConstr.of_constr typ)) + try find_clause (EConstr.of_constr (try_red_product env sigma typ)) with Redelimination -> raise e in find_clause typ -- cgit v1.2.3 From 7b43de20a4acd7c9da290f038d9a16fe67eccd59 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sat, 19 Nov 2016 01:59:07 +0100 Subject: Leminv API using EConstr. --- engine/evarutil.ml | 8 ++++++-- engine/evarutil.mli | 4 ++-- engine/termops.ml | 4 ++-- engine/termops.mli | 4 ++-- pretyping/cases.ml | 14 ++++--------- pretyping/inductiveops.ml | 14 +++++++------ pretyping/inductiveops.mli | 12 +++++------ pretyping/pretyping.ml | 1 - pretyping/retyping.ml | 2 +- tactics/inv.ml | 1 - tactics/leminv.ml | 50 ++++++++++++++++++++++++++++++---------------- tactics/tactics.ml | 9 ++++----- toplevel/command.ml | 4 ++-- 13 files changed, 70 insertions(+), 57 deletions(-) diff --git a/engine/evarutil.ml b/engine/evarutil.ml index c2ad3c462..35fe9cf66 100644 --- a/engine/evarutil.ml +++ b/engine/evarutil.ml @@ -68,6 +68,9 @@ 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 flush_and_check_evars sigma c = + flush_and_check_evars sigma (EConstr.Unsafe.to_constr c) + (* let nf_evar_key = Profile.declare_profile "nf_evar" *) (* let nf_evar = Profile.profile2 nf_evar_key Reductionops.nf_evar *) @@ -474,12 +477,13 @@ let e_new_evar env evdref ?(src=default_source) ?filter ?candidates ?store ?nami (* This assumes an evar with identity instance and generalizes it over only 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 *) diff --git a/engine/evarutil.mli b/engine/evarutil.mli index 82346b24e..ad3851ea3 100644 --- a/engine/evarutil.mli +++ b/engine/evarutil.mli @@ -170,7 +170,7 @@ val nf_evar_map_universes : evar_map -> evar_map * (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 -> EConstr.constr -> constr (** {6 Term manipulation up to instantiation} *) @@ -220,7 +220,7 @@ val push_rel_decl_to_named_context : val push_rel_context_to_named_context : Environ.env -> EConstr.types -> named_context_val * EConstr.types * EConstr.constr list * csubst * (identifier*EConstr.constr) list -val generalize_evar_over_rels : evar_map -> existential -> types * constr list +val generalize_evar_over_rels : evar_map -> EConstr.existential -> EConstr.types * EConstr.constr list (** Evar combinators *) diff --git a/engine/termops.ml b/engine/termops.ml index 4ab9f0733..ef7cdc38b 100644 --- a/engine/termops.ml +++ b/engine/termops.ml @@ -233,9 +233,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 diff --git a/engine/termops.mli b/engine/termops.mli index 1c14e6c89..72c0cedda 100644 --- a/engine/termops.mli +++ b/engine/termops.mli @@ -51,9 +51,9 @@ val it_mkLambda : EConstr.constr -> (Name.t * EConstr.types) list -> EConstr.con val it_mkProd_or_LetIn : types -> Context.Rel.t -> types val it_mkProd_wo_LetIn : EConstr.types -> Context.Rel.t -> EConstr.types val it_mkLambda_or_LetIn : constr -> Context.Rel.t -> constr -val it_mkNamedProd_or_LetIn : types -> Context.Named.t -> types +val it_mkNamedProd_or_LetIn : EConstr.types -> Context.Named.t -> EConstr.types val it_mkNamedProd_wo_LetIn : types -> Context.Named.t -> types -val it_mkNamedLambda_or_LetIn : constr -> Context.Named.t -> constr +val it_mkNamedLambda_or_LetIn : EConstr.constr -> Context.Named.t -> EConstr.constr (* Ad hoc version reinserting letin, assuming the body is defined in the context where the letins are expanded *) diff --git a/pretyping/cases.ml b/pretyping/cases.ml index 360199fec..119e92c82 100644 --- a/pretyping/cases.ml +++ b/pretyping/cases.ml @@ -326,7 +326,7 @@ let inh_coerce_to_ind evdref env loc ty tyi = let binding_vars_of_inductive sigma = function | NotInd _ -> [] - | IsInd (_,IndType(_,realargs),_) -> List.filter (isRel sigma) (List.map EConstr.of_constr realargs) + | IsInd (_,IndType(_,realargs),_) -> List.filter (isRel sigma) realargs let extract_inductive_data env sigma decl = match decl with @@ -422,7 +422,7 @@ let type_of_tomatch = function | NotInd (_,t) -> t let map_tomatch_type f = function - | IsInd (t,ind,names) -> IsInd (f t,map_inductive_type (fun c -> EConstr.Unsafe.to_constr (f (EConstr.of_constr c))) ind,names) + | IsInd (t,ind,names) -> IsInd (f t,map_inductive_type f ind,names) | NotInd (c,t) -> NotInd (Option.map f c, f t) let liftn_tomatch_type n depth = map_tomatch_type (Vars.liftn n depth) @@ -873,7 +873,7 @@ let specialize_predicate_var (cur,typ,dep) tms ccl = let l = match typ with | IsInd (_, IndType (_, _), []) -> [] - | IsInd (_, IndType (_, realargs), names) -> List.map EConstr.of_constr realargs + | IsInd (_, IndType (_, realargs), names) -> realargs | NotInd _ -> [] in subst_predicate (l,c) ccl tms @@ -922,7 +922,7 @@ let rec extract_predicate ccl = function subst1 cur pred end | Pushed (_,((cur,IsInd (_,IndType(_,realargs),_)),_,na))::tms -> - let realargs = List.rev_map EConstr.of_constr realargs in + let realargs = List.rev realargs in let k, nrealargs = match na with | Anonymous -> 0, realargs | Name _ -> 1, (cur :: realargs) @@ -1064,7 +1064,6 @@ let specialize_predicate newtomatchs (names,depna) arsign cs tms ccl = snd (List.fold_left (expand_arg tms) (1,ccl''') newtomatchs) let find_predicate loc env evdref p current (IndType (indf,realargs)) dep tms = - let realargs = List.map EConstr.of_constr realargs in let pred = abstract_predicate env !evdref indf current realargs dep tms p in (pred, EConstr.of_constr (whd_betaiota !evdref (applist (pred, realargs@[current])))) @@ -1384,7 +1383,6 @@ and match_current pb (initial,tomatch) = if (not no_cstr || not (List.is_empty pb.mat)) && onlydflt then compile_all_variables initial tomatch pb else - let realargs = List.map EConstr.of_constr realargs in (* We generalize over terms depending on current term to match *) let pb,deps = generalize_problem (names,dep) pb deps in @@ -1749,7 +1747,6 @@ let build_inversion_problem loc env sigma tms t = match tms with | [] -> [], acc_sign, acc | (t, IsInd (_,IndType(indf,realargs),_)) :: tms -> - let realargs = List.map EConstr.of_constr realargs in let patl,acc = List.fold_map' reveal_pattern realargs acc in let pat,acc = make_patvar t acc in let indf' = lift_inductive_family n indf in @@ -1919,7 +1916,6 @@ let prepare_predicate_from_arsign_tycon env sigma loc tomatchs arsign c = (match tmtype with NotInd _ -> (subst, len - signlen) | IsInd (_, IndType(indf,realargs),_) -> - let realargs = List.map EConstr.of_constr realargs in let subst, len = List.fold_left (fun (subst, len) arg -> @@ -2119,7 +2115,6 @@ let constr_of_pat env evdref arsign pat avoid = let apptype = Retyping.get_type_of env ( !evdref) app in let apptype = EConstr.of_constr apptype in let IndType (indf, realargs) = find_rectype env (!evdref) apptype in - let realargs = List.map EConstr.of_constr realargs in match alias with Anonymous -> pat', sign, app, apptype, realargs, n, avoid @@ -2364,7 +2359,6 @@ let build_dependent_signature env evdref avoid tomatchs arsign = *) match ty with | IsInd (ty, IndType (indf, args), _) when List.length args > 0 -> - let args = List.map EConstr.of_constr args in (* Build the arity signature following the names in matched terms as much as possible *) let argsign = List.tl arsign in (* arguments in inverse application order *) diff --git a/pretyping/inductiveops.ml b/pretyping/inductiveops.ml index 98b267cfd..cb8b25323 100644 --- a/pretyping/inductiveops.ml +++ b/pretyping/inductiveops.ml @@ -58,21 +58,23 @@ let lift_inductive_family n = liftn_inductive_family n 1 let substnl_ind_family l n = map_ind_family (substnl l n) -type inductive_type = IndType of inductive_family * constr list +type inductive_type = IndType of inductive_family * EConstr.constr list let make_ind_type (indf, realargs) = IndType (indf,realargs) let dest_ind_type (IndType (indf,realargs)) = (indf,realargs) let map_inductive_type f (IndType (indf, realargs)) = - IndType (map_ind_family f indf, List.map f realargs) + let f' c = EConstr.Unsafe.to_constr (f (EConstr.of_constr c)) in + IndType (map_ind_family f' indf, List.map f realargs) -let liftn_inductive_type n d = map_inductive_type (liftn n d) +let liftn_inductive_type n d = map_inductive_type (EConstr.Vars.liftn n d) let lift_inductive_type n = liftn_inductive_type n 1 -let substnl_ind_type l n = map_inductive_type (substnl l n) +let substnl_ind_type l n = map_inductive_type (EConstr.Vars.substnl l n) let mkAppliedInd (IndType ((ind,params), realargs)) = - applist (mkIndU ind,params@realargs) + let open EConstr in + applist (mkIndU ind, (List.map EConstr.of_constr params)@realargs) (* Does not consider imbricated or mutually recursive types *) let mis_is_recursive_subset listind rarg = @@ -471,7 +473,7 @@ let find_rectype env sigma c = if mib.mind_nparams > List.length l then raise Not_found; let l = List.map EConstr.Unsafe.to_constr l in let (par,rargs) = List.chop mib.mind_nparams l in - IndType((indu, par),rargs) + IndType((indu, par),List.map EConstr.of_constr rargs) | _ -> raise Not_found let find_inductive env sigma c = diff --git a/pretyping/inductiveops.mli b/pretyping/inductiveops.mli index 7af9731f9..1614e1817 100644 --- a/pretyping/inductiveops.mli +++ b/pretyping/inductiveops.mli @@ -37,15 +37,15 @@ val substnl_ind_family : constr list -> int -> inductive_family -> inductive_family (** An inductive type with its parameters and real arguments *) -type inductive_type = IndType of inductive_family * constr list -val make_ind_type : inductive_family * constr list -> inductive_type -val dest_ind_type : inductive_type -> inductive_family * constr list -val map_inductive_type : (constr -> constr) -> inductive_type -> inductive_type +type inductive_type = IndType of inductive_family * EConstr.constr list +val make_ind_type : inductive_family * EConstr.constr list -> inductive_type +val dest_ind_type : inductive_type -> inductive_family * EConstr.constr list +val map_inductive_type : (EConstr.constr -> EConstr.constr) -> inductive_type -> inductive_type val liftn_inductive_type : int -> int -> inductive_type -> inductive_type val lift_inductive_type : int -> inductive_type -> inductive_type -val substnl_ind_type : constr list -> int -> inductive_type -> inductive_type +val substnl_ind_type : EConstr.constr list -> int -> inductive_type -> inductive_type -val mkAppliedInd : inductive_type -> constr +val mkAppliedInd : inductive_type -> EConstr.constr val mis_is_recursive_subset : int list -> wf_paths -> bool val mis_is_recursive : inductive * mutual_inductive_body * one_inductive_body -> bool diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml index 7586b54ba..11d50926f 100644 --- a/pretyping/pretyping.ml +++ b/pretyping/pretyping.ml @@ -860,7 +860,6 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre let cloc = loc_of_glob_constr c in error_case_not_inductive ~loc:cloc env.ExtraEnv.env !evdref cj in - let realargs = List.map EConstr.of_constr realargs in let cstrs = get_constructors env.ExtraEnv.env indf in if not (Int.equal (Array.length cstrs) 1) then user_err ~loc (str "Destructing let is only for inductive types" ++ diff --git a/pretyping/retyping.ml b/pretyping/retyping.ml index 6f03fc462..88899e633 100644 --- a/pretyping/retyping.ml +++ b/pretyping/retyping.ml @@ -123,7 +123,7 @@ let retype ?(polyprop=true) sigma = with Not_found -> retype_error BadRecursiveType in let n = inductive_nrealdecls_env env (fst (fst (dest_ind_family indf))) in - let t = EConstr.of_constr (betazetaevar_applist sigma n p (List.map EConstr.of_constr realargs)) in + let t = EConstr.of_constr (betazetaevar_applist sigma n p realargs) in (match EConstr.kind sigma (EConstr.of_constr (whd_all env sigma (type_of env t))) with | Prod _ -> EConstr.of_constr (whd_beta sigma (applist (t, [c]))) | _ -> t) diff --git a/tactics/inv.ml b/tactics/inv.ml index 37c82ff64..5c296b343 100644 --- a/tactics/inv.ml +++ b/tactics/inv.ml @@ -450,7 +450,6 @@ let raw_inversion inv_kind id status names = CErrors.user_err msg in let IndType (indf,realargs) = find_rectype env sigma t in - let realargs = List.map EConstr.of_constr realargs in let evdref = ref sigma in let (elim_predicate, args) = make_inv_predicate env evdref indf realargs id status concl in diff --git a/tactics/leminv.ml b/tactics/leminv.ml index a94238418..62e78b588 100644 --- a/tactics/leminv.ml +++ b/tactics/leminv.ml @@ -11,8 +11,9 @@ open CErrors open Util open Names open Term -open Vars open Termops +open EConstr +open Vars open Namegen open Evd open Printer @@ -31,9 +32,17 @@ open Context.Named.Declaration module NamedDecl = Context.Named.Declaration +let nlocal_assum (na, t) = + let inj = EConstr.Unsafe.to_constr in + NamedDecl.LocalAssum (na, inj t) + +let nlocal_def (na, b, t) = + let inj = EConstr.Unsafe.to_constr in + NamedDecl.LocalDef (na, inj b, inj t) + let no_inductive_inconstr env sigma constr = (str "Cannot recognize an inductive predicate in " ++ - pr_lconstr_env env sigma constr ++ + pr_lconstr_env env sigma (EConstr.Unsafe.to_constr constr) ++ str "." ++ spc () ++ str "If there is one, may be the structure of the arity" ++ spc () ++ str "or of the type of constructors" ++ spc () ++ str "is hidden by constant definitions.") @@ -116,15 +125,15 @@ let max_prefix_sign lid sign = | id::l -> snd (max_rec (id, sign_prefix id sign) l) *) let rec add_prods_sign env sigma t = - match kind_of_term (whd_all env sigma (EConstr.of_constr t)) with + match EConstr.kind sigma (EConstr.of_constr (whd_all env sigma t)) with | Prod (na,c1,b) -> - let id = id_of_name_using_hdchar env t na in + let id = id_of_name_using_hdchar env (EConstr.Unsafe.to_constr t) na in let b'= subst1 (mkVar id) b in - add_prods_sign (push_named (LocalAssum (id,c1)) env) sigma b' + add_prods_sign (push_named (nlocal_assum (id,c1)) env) sigma b' | LetIn (na,c1,t1,b) -> - let id = id_of_name_using_hdchar env t na in + let id = id_of_name_using_hdchar env (EConstr.Unsafe.to_constr t) na in let b'= subst1 (mkVar id) b in - add_prods_sign (push_named (LocalDef (id,c1,t1)) env) sigma b' + add_prods_sign (push_named (nlocal_def (id,c1,t1)) env) sigma b' | _ -> (env,t) (* [dep_option] indicates whether the inversion lemma is dependent or not. @@ -147,6 +156,7 @@ let compute_first_inversion_scheme env sigma ind sort dep_option = let pty,goal = if dep_option then let pty = make_arity env true indf sort in + let pty = EConstr.of_constr pty in let goal = mkProd (Anonymous, mkAppliedInd ind, applist(mkVar p,realargs@[mkRel 1])) @@ -154,7 +164,7 @@ let compute_first_inversion_scheme env sigma ind sort dep_option = pty,goal else let i = mkAppliedInd ind in - let ivars = global_vars env sigma (EConstr.of_constr i) in + let ivars = global_vars env sigma i in let revargs,ownsign = fold_named_context (fun env d (revargs,hyps) -> @@ -169,7 +179,7 @@ let compute_first_inversion_scheme env sigma ind sort dep_option = let goal = mkArrow i (applist(mkVar p, List.rev revargs)) in (pty,goal) in - let npty = nf_all env sigma (EConstr.of_constr pty) in + let npty = nf_all env sigma pty in let extenv = push_named (LocalAssum (p,npty)) env in extenv, goal @@ -183,7 +193,7 @@ let compute_first_inversion_scheme env sigma ind sort dep_option = let inversion_scheme env sigma t sort dep_option inv_op = let (env,i) = add_prods_sign env sigma t in let ind = - try find_rectype env sigma (EConstr.of_constr i) + try find_rectype env sigma i with Not_found -> user_err ~hdr:"inversion_scheme" (no_inductive_inconstr env sigma i) in @@ -192,18 +202,20 @@ let inversion_scheme env sigma t sort dep_option inv_op = in assert (List.subset - (global_vars env sigma (EConstr.of_constr invGoal)) + (global_vars env sigma invGoal) (ids_of_named_context (named_context invEnv))); (* user_err ~hdr:"lemma_inversion" (str"Computed inversion goal was not closed in initial signature."); *) + let invGoal = EConstr.Unsafe.to_constr invGoal in let pf = Proof.start (Evd.from_ctx (evar_universe_context sigma)) [invEnv,invGoal] in let pf = fst (Proof.run_tactic env ( tclTHEN intro (onLastHypId inv_op)) pf) in let pfterm = List.hd (Proof.partial_proof pf) in + let pfterm = EConstr.of_constr pfterm in let global_named_context = Global.named_context_val () in let ownSign = ref begin fold_named_context @@ -216,18 +228,19 @@ let inversion_scheme env sigma t sort dep_option inv_op = let { sigma=sigma } = Proof.V82.subgoals pf in let sigma = Evd.nf_constraints sigma in let rec fill_holes c = - match kind_of_term c with + match EConstr.kind sigma c with | Evar (e,args) -> let h = next_ident_away (Id.of_string "H") !avoid in let ty,inst = Evarutil.generalize_evar_over_rels sigma (e,args) in avoid := h::!avoid; - ownSign := Context.Named.add (LocalAssum (h,ty)) !ownSign; + ownSign := Context.Named.add (nlocal_assum (h,ty)) !ownSign; applist (mkVar h, inst) - | _ -> Constr.map fill_holes c + | _ -> EConstr.map sigma fill_holes c in let c = fill_holes pfterm in (* warning: side-effect on ownSign *) let invProof = it_mkNamedLambda_or_LetIn c !ownSign in + let invProof = EConstr.Unsafe.to_constr invProof in let p = Evarutil.nf_evars_universes sigma invProof in p, Evd.universe_context sigma @@ -245,6 +258,7 @@ let add_inversion_lemma_exn na com comsort bool tac = let env = Global.env () in let evd = ref (Evd.from_env env) in let c = Constrintern.interp_type_evars env evd com in + let c = EConstr.of_constr c in let sigma, sort = Pretyping.interp_sort !evd comsort in try add_inversion_lemma na env sigma c sort bool tac @@ -258,14 +272,16 @@ let add_inversion_lemma_exn na com comsort bool tac = let lemInv id c gls = try - let clause = mk_clenv_type_of gls (EConstr.of_constr c) in + let clause = mk_clenv_type_of gls c in let clause = clenv_constrain_last_binding (EConstr.mkVar id) clause in Proofview.V82.of_tactic (Clenvtac.res_pf clause ~flags:(Unification.elim_flags ()) ~with_evars:false) gls with | NoSuchBinding -> + let c = EConstr.Unsafe.to_constr c in user_err (hov 0 (pr_constr c ++ spc () ++ str "does not refer to an inversion lemma.")) | UserError (a,b) -> + let c = EConstr.Unsafe.to_constr c in user_err ~hdr:"LemInv" (str "Cannot refine current goal with the lemma " ++ pr_lconstr_env (Refiner.pf_env gls) (Refiner.project gls) c) @@ -291,5 +307,5 @@ let lemInvIn id c ids = let lemInvIn_gen id c l = try_intros_until (fun id -> lemInvIn id c l) id let lemInv_clause id c = function - | [] -> lemInv_gen id (EConstr.Unsafe.to_constr c) - | l -> lemInvIn_gen id (EConstr.Unsafe.to_constr c) l + | [] -> lemInv_gen id c + | l -> lemInvIn_gen id c l diff --git a/tactics/tactics.ml b/tactics/tactics.ml index bae1ad48c..59ffd8b62 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -2896,8 +2896,7 @@ let old_generalize_dep ?(with_let=false) c gl = -> id::tothin | _ -> tothin in - let cl' = it_mkNamedProd_or_LetIn (Tacmach.pf_concl gl) to_quantify in - let cl' = EConstr.of_constr cl' in + let cl' = it_mkNamedProd_or_LetIn (EConstr.of_constr (Tacmach.pf_concl gl)) to_quantify in let body = if with_let then match EConstr.kind sigma c with @@ -4256,11 +4255,11 @@ let apply_induction_in_context with_evars hyp0 inhyps elim indvars names induct_ let env = Proofview.Goal.env gl in let sigma = Sigma.to_evar_map sigma in let concl = Tacmach.New.pf_nf_concl gl in + let concl = EConstr.of_constr concl in let statuslists,lhyp0,toclear,deps,avoid,dep_in_hyps = cook_sign hyp0 inhyps indvars env sigma in - let dep_in_concl = Option.cata (fun id -> occur_var env sigma id (EConstr.of_constr concl)) false hyp0 in + let dep_in_concl = Option.cata (fun id -> occur_var env sigma id concl) false hyp0 in let dep = dep_in_hyps || dep_in_concl in let tmpcl = it_mkNamedProd_or_LetIn concl deps in - let tmpcl = EConstr.of_constr tmpcl in let s = Retyping.get_sort_family_of env sigma tmpcl in let deps_cstr = List.fold_left @@ -5008,7 +5007,7 @@ let abstract_subproof id gk tac = else (Context.Named.add d s1,s2)) global_sign (Context.Named.empty, empty_named_context_val) in let id = next_global_ident_away id (pf_ids_of_hyps gl) in - let concl = it_mkNamedProd_or_LetIn (Proofview.Goal.concl gl) sign in + let concl = it_mkNamedProd_or_LetIn (EConstr.of_constr (Proofview.Goal.concl gl)) sign in let concl = try flush_and_check_evars !evdref concl with Uninstantiated_evar _ -> diff --git a/toplevel/command.ml b/toplevel/command.ml index 08f3ad4a7..1b5c4f660 100644 --- a/toplevel/command.ml +++ b/toplevel/command.ml @@ -1271,9 +1271,9 @@ let do_program_recursive local p fixkind fixl ntns = let collect_evars id def typ imps = (* Generalize by the recursive prototypes *) let def = - nf_evar evd (Termops.it_mkNamedLambda_or_LetIn def rec_sign) + nf_evar evd (EConstr.Unsafe.to_constr (Termops.it_mkNamedLambda_or_LetIn (EConstr.of_constr def) rec_sign)) and typ = - nf_evar evd (Termops.it_mkNamedProd_or_LetIn typ rec_sign) + nf_evar evd (EConstr.Unsafe.to_constr (Termops.it_mkNamedProd_or_LetIn (EConstr.of_constr typ) rec_sign)) in let evm = collect_evars_of_term evd def typ in let evars, _, def, typ = -- cgit v1.2.3 From 34e86e839be251717db96f1f5969d7724ab43097 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sat, 19 Nov 2016 02:45:54 +0100 Subject: Hints API using EConstr. --- engine/evarutil.ml | 6 +- engine/evarutil.mli | 2 +- engine/termops.ml | 5 +- engine/termops.mli | 2 +- ltac/g_auto.ml4 | 2 +- ltac/g_class.ml4 | 2 +- plugins/firstorder/sequent.ml | 4 +- pretyping/evarsolve.ml | 10 ++- proofs/evar_refiner.ml | 10 ++- tactics/auto.ml | 33 ++++---- tactics/auto.mli | 4 +- tactics/class_tactics.ml | 51 ++++++------- tactics/class_tactics.mli | 2 +- tactics/eauto.ml | 32 ++++---- tactics/eauto.mli | 2 +- tactics/hints.ml | 173 +++++++++++++++++++++++------------------- tactics/hints.mli | 9 ++- toplevel/classes.ml | 2 +- 18 files changed, 188 insertions(+), 163 deletions(-) diff --git a/engine/evarutil.ml b/engine/evarutil.ml index 35fe9cf66..8b75d8242 100644 --- a/engine/evarutil.ml +++ b/engine/evarutil.ml @@ -187,7 +187,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 @@ -196,7 +198,7 @@ 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) *) diff --git a/engine/evarutil.mli b/engine/evarutil.mli index ad3851ea3..1b592b790 100644 --- a/engine/evarutil.mli +++ b/engine/evarutil.mli @@ -88,7 +88,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 -> EConstr.constr -> existential_key (** may raise NoHeadEvar *) (* Expand head evar if any *) val whd_head_evar : evar_map -> EConstr.constr -> EConstr.constr diff --git a/engine/termops.ml b/engine/termops.ml index ef7cdc38b..7c89f190f 100644 --- a/engine/termops.ml +++ b/engine/termops.ml @@ -263,9 +263,10 @@ 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 EConstr.isEvar sigma (Array.last args) -> + let open EConstr in drop_extra_implicit_args sigma - (EConstr.mkApp (f,fst (Array.chop (Array.length args - 1) args))) - | _ -> EConstr.Unsafe.to_constr c + (mkApp (f,fst (Array.chop (Array.length args - 1) args))) + | _ -> c (* Get the last arg of an application *) let last_arg sigma c = match EConstr.kind sigma c with diff --git a/engine/termops.mli b/engine/termops.mli index 72c0cedda..a865c80fb 100644 --- a/engine/termops.mli +++ b/engine/termops.mli @@ -96,7 +96,7 @@ val iter_constr_with_full_binders : (**********************************************************************) val strip_head_cast : Evd.evar_map -> EConstr.t -> EConstr.t -val drop_extra_implicit_args : Evd.evar_map -> EConstr.t -> constr +val drop_extra_implicit_args : Evd.evar_map -> EConstr.t -> EConstr.constr (** occur checks *) diff --git a/ltac/g_auto.ml4 b/ltac/g_auto.ml4 index c6395d7e2..f540beb56 100644 --- a/ltac/g_auto.ml4 +++ b/ltac/g_auto.ml4 @@ -27,7 +27,7 @@ TACTIC EXTEND eassumption END TACTIC EXTEND eexact -| [ "eexact" constr(c) ] -> [ Eauto.e_give_exact c ] +| [ "eexact" constr(c) ] -> [ Eauto.e_give_exact (EConstr.of_constr c) ] END let pr_hintbases _prc _prlc _prt = Pptactic.pr_hintbases diff --git a/ltac/g_class.ml4 b/ltac/g_class.ml4 index f8654d390..ea9a2b6e1 100644 --- a/ltac/g_class.ml4 +++ b/ltac/g_class.ml4 @@ -74,7 +74,7 @@ TACTIC EXTEND is_ground END TACTIC EXTEND autoapply - [ "autoapply" constr(c) "using" preident(i) ] -> [ Proofview.V82.tactic (autoapply c i) ] + [ "autoapply" constr(c) "using" preident(i) ] -> [ Proofview.V82.tactic (autoapply (EConstr.of_constr c) i) ] END (** TODO: DEPRECATE *) diff --git a/plugins/firstorder/sequent.ml b/plugins/firstorder/sequent.ml index 87e7192d7..91cd102a2 100644 --- a/plugins/firstorder/sequent.ml +++ b/plugins/firstorder/sequent.ml @@ -214,8 +214,8 @@ let extend_with_auto_hints l seq gl= | Res_pf_THEN_trivial_fail (c,_) -> let (c, _, _) = c in (try - let gr = global_of_constr c in - let typ=(pf_unsafe_type_of gl (EConstr.of_constr c)) in + let (gr, _) = Termops.global_of_constr (project gl) c in + let typ=(pf_unsafe_type_of gl c) in seqref:=add_formula Hint gr typ !seqref gl with Not_found->()) | _-> () in diff --git a/pretyping/evarsolve.ml b/pretyping/evarsolve.ml index 4ce8a44ad..772571926 100644 --- a/pretyping/evarsolve.ml +++ b/pretyping/evarsolve.ml @@ -1612,12 +1612,14 @@ and evar_define conv_algo ?(choose=false) env evd pbty (evk,argsv as ev) rhs = * ass. *) -let status_changed lev (pbty,_,t1,t2) = - (try Evar.Set.mem (head_evar t1) lev with NoHeadEvar -> false) || - (try Evar.Set.mem (head_evar t2) lev with NoHeadEvar -> false) +let status_changed evd lev (pbty,_,t1,t2) = + let t1 = EConstr.of_constr t1 in + let t2 = EConstr.of_constr t2 in + (try Evar.Set.mem (head_evar evd t1) lev with NoHeadEvar -> false) || + (try Evar.Set.mem (head_evar evd t2) lev with NoHeadEvar -> false) let reconsider_conv_pbs conv_algo evd = - let (evd,pbs) = extract_changed_conv_pbs evd status_changed in + let (evd,pbs) = extract_changed_conv_pbs evd (status_changed evd) in List.fold_left (fun p (pbty,env,t1,t2 as x) -> match p with diff --git a/proofs/evar_refiner.ml b/proofs/evar_refiner.ml index 4be03af9a..0d65faf12 100644 --- a/proofs/evar_refiner.ml +++ b/proofs/evar_refiner.ml @@ -18,17 +18,19 @@ open Pp (* Instantiation of existential variables *) (******************************************) -let depends_on_evar evk _ (pbty,_,t1,t2) = - try Evar.equal (head_evar t1) evk +let depends_on_evar sigma evk _ (pbty,_,t1,t2) = + let t1 = EConstr.of_constr t1 in + let t2 = EConstr.of_constr t2 in + try Evar.equal (head_evar sigma t1) evk with NoHeadEvar -> - try Evar.equal (head_evar t2) evk + try Evar.equal (head_evar sigma t2) evk with NoHeadEvar -> false let define_and_solve_constraints evk c env evd = if Termops.occur_evar evd evk (EConstr.of_constr c) then Pretype_errors.error_occur_check env evd evk (EConstr.of_constr c); let evd = define evk c evd in - let (evd,pbs) = extract_changed_conv_pbs evd (depends_on_evar evk) in + let (evd,pbs) = extract_changed_conv_pbs evd (depends_on_evar evd evk) in match List.fold_left (fun p (pbty,env,t1,t2) -> match p with diff --git a/tactics/auto.ml b/tactics/auto.ml index 41b56bd3d..2423ea878 100644 --- a/tactics/auto.ml +++ b/tactics/auto.ml @@ -93,7 +93,7 @@ let connect_hint_clenv poly (c, _, ctx) clenv gl = evd = Evd.map_metas map evd; env = Proofview.Goal.env gl; } in - clenv, map c + clenv, emap c else let evd = Evd.merge_context_set Evd.univ_flexible evd ctx in { clenv with evd = evd ; env = Proofview.Goal.env gl }, c @@ -115,7 +115,6 @@ let unify_resolve_gen poly = function let exact poly (c,clenv) = Proofview.Goal.enter { enter = begin fun gl -> let clenv', c = connect_hint_clenv poly c clenv gl in - let c = EConstr.of_constr c in Tacticals.New.tclTHEN (Proofview.Unsafe.tclEVARUNIVCONTEXT (Evd.evar_universe_context clenv'.evd)) (exact_check c) @@ -141,7 +140,7 @@ let conclPattern concl pat tac = | None -> Proofview.tclUNIT Id.Map.empty | Some pat -> try - Proofview.tclUNIT (Constr_matching.matches env sigma pat (EConstr.of_constr concl)) + Proofview.tclUNIT (Constr_matching.matches env sigma pat concl) with Constr_matching.PatternMatchingFailure -> Tacticals.New.tclZEROMSG (str "conclPattern") in @@ -300,13 +299,13 @@ let flags_of_state st = let auto_flags_of_state st = auto_unif_flags_of full_transparent_state st false -let hintmap_of secvars hdc concl = +let hintmap_of sigma secvars hdc concl = match hdc with | None -> Hint_db.map_none ~secvars | Some hdc -> - if occur_existential Evd.empty (EConstr.of_constr concl) then (** FIXME *) - Hint_db.map_existential ~secvars hdc concl - else Hint_db.map_auto ~secvars hdc concl + if occur_existential sigma concl then + Hint_db.map_existential sigma ~secvars hdc concl + else Hint_db.map_auto sigma ~secvars hdc concl let exists_evaluable_reference env = function | EvalConstRef _ -> true @@ -331,6 +330,7 @@ let rec trivial_fail_db dbg mod_delta db_list local_db = in Proofview.Goal.enter { enter = begin fun gl -> let concl = Tacmach.New.pf_nf_concl gl in + let concl = EConstr.of_constr concl in let sigma = Tacmach.New.project gl in let secvars = compute_secvars gl in Tacticals.New.tclFIRST @@ -339,17 +339,17 @@ let rec trivial_fail_db dbg mod_delta db_list local_db = (trivial_resolve sigma dbg mod_delta db_list local_db secvars concl))) end } -and my_find_search_nodelta db_list local_db secvars hdc concl = +and my_find_search_nodelta sigma db_list local_db secvars hdc concl = List.map (fun hint -> (None,hint)) - (List.map_append (hintmap_of secvars hdc concl) (local_db::db_list)) + (List.map_append (hintmap_of sigma secvars hdc concl) (local_db::db_list)) and my_find_search mod_delta = if mod_delta then my_find_search_delta else my_find_search_nodelta -and my_find_search_delta db_list local_db secvars hdc concl = - let f = hintmap_of secvars hdc concl in - if occur_existential Evd.empty (EConstr.of_constr concl) (** FIXME *) then +and my_find_search_delta sigma db_list local_db secvars hdc concl = + let f = hintmap_of sigma secvars hdc concl in + if occur_existential sigma concl then List.map_append (fun db -> if Hint_db.use_dn db then @@ -371,8 +371,8 @@ and my_find_search_delta db_list local_db secvars hdc concl = match hdc with None -> Hint_db.map_none ~secvars db | Some hdc -> if (Id.Pred.is_empty ids && Cpred.is_empty csts) - then Hint_db.map_auto ~secvars hdc concl db - else Hint_db.map_existential ~secvars hdc concl db + then Hint_db.map_auto sigma ~secvars hdc concl db + else Hint_db.map_existential sigma ~secvars hdc concl db in auto_flags_of_state st, l in List.map (fun x -> (Some flags,x)) l) (local_db::db_list) @@ -414,7 +414,7 @@ and trivial_resolve sigma dbg mod_delta db_list local_db secvars cl = in List.map (tac_of_hint dbg db_list local_db cl) (priority - (my_find_search mod_delta db_list local_db secvars head cl)) + (my_find_search mod_delta sigma db_list local_db secvars head cl)) with Not_found -> [] (** The use of the "core" database can be de-activated by passing @@ -460,7 +460,7 @@ let possible_resolve sigma dbg mod_delta db_list local_db secvars cl = with Bound -> None in List.map (tac_of_hint dbg db_list local_db cl) - (my_find_search mod_delta db_list local_db secvars head cl) + (my_find_search mod_delta sigma db_list local_db secvars head cl) with Not_found -> [] let extend_local_db decl db gl = @@ -491,6 +491,7 @@ let search d n mod_delta db_list local_db = (Tacticals.New.tclORELSE0 (intro_register d (search d n) local_db) ( Proofview.Goal.enter { enter = begin fun gl -> let concl = Tacmach.New.pf_nf_concl gl in + let concl = EConstr.of_constr concl in let sigma = Tacmach.New.project gl in let secvars = compute_secvars gl in let d' = incr_dbg d in diff --git a/tactics/auto.mli b/tactics/auto.mli index 06048ac1c..403be9e1c 100644 --- a/tactics/auto.mli +++ b/tactics/auto.mli @@ -23,7 +23,7 @@ val default_search_depth : int ref val auto_flags_of_state : transparent_state -> Unification.unify_flags val connect_hint_clenv : polymorphic -> raw_hint -> clausenv -> - ('a, 'r) Proofview.Goal.t -> clausenv * constr + ('a, 'r) Proofview.Goal.t -> clausenv * EConstr.constr (** Try unification with the precompiled clause, then use registered Apply *) val unify_resolve : polymorphic -> Unification.unify_flags -> (raw_hint * clausenv) -> unit Proofview.tactic @@ -33,7 +33,7 @@ val unify_resolve : polymorphic -> Unification.unify_flags -> (raw_hint * clause [Pattern.somatches], then replace [?1] [?2] metavars in tacast by the right values to build a tactic *) -val conclPattern : constr -> constr_pattern option -> Genarg.glob_generic_argument -> unit Proofview.tactic +val conclPattern : EConstr.constr -> constr_pattern option -> Genarg.glob_generic_argument -> unit Proofview.tactic (** The Auto tactic *) diff --git a/tactics/class_tactics.ml b/tactics/class_tactics.ml index 02211efd6..8ecdd01f2 100644 --- a/tactics/class_tactics.ml +++ b/tactics/class_tactics.ml @@ -217,6 +217,7 @@ let auto_unif_flags freeze st = } let e_give_exact flags poly (c,clenv) gl = + let open EConstr in let (c, _, _) = c in let c, gl = if poly then @@ -226,7 +227,6 @@ let e_give_exact flags poly (c,clenv) gl = c, {gl with sigma = evd} else c, gl in - let c = EConstr.of_constr c in let t1 = pf_unsafe_type_of gl c in let t1 = EConstr.of_constr t1 in Proofview.V82.of_tactic (Clenvtac.unify ~flags t1 <*> exact_no_check c) gl @@ -245,6 +245,7 @@ let unify_resolve poly flags = { enter = begin fun gls (c,_,clenv) -> (** Application of a lemma using [refine] instead of the old [w_unify] *) let unify_resolve_refine poly flags = + let open EConstr in let open Clenv in { enter = begin fun gls ((c, t, ctx),n,clenv) -> let env = Proofview.Goal.env gls in @@ -262,9 +263,6 @@ let unify_resolve_refine poly flags = let sigma = Evd.merge_context_set Evd.univ_flexible sigma ctx in sigma, c, t in - let open EConstr in - let ty = EConstr.of_constr ty in - let term = EConstr.of_constr term in let sigma', cl = Clenv.make_evar_clause env sigma ?len:n ty in let term = applist (term, List.map (fun x -> x.hole_evar) cl.cl_holes) in let sigma' = @@ -285,7 +283,6 @@ let clenv_of_prods poly nprods (c, clenv) gl = let (c, _, _) = c in if poly || Int.equal nprods 0 then Some (None, clenv) else - let c = EConstr.of_constr c in let sigma = Tacmach.New.project gl in let ty = Retyping.get_type_of (Proofview.Goal.env gl) sigma c in let ty = EConstr.of_constr ty in @@ -313,7 +310,7 @@ let matches_pattern concl pat = | None -> Proofview.tclUNIT () | Some pat -> let sigma = Sigma.to_evar_map sigma in - if Constr_matching.is_matching env sigma pat (EConstr.of_constr concl) then + if Constr_matching.is_matching env sigma pat concl then Proofview.tclUNIT () else Tacticals.New.tclZEROMSG (str "conclPattern") @@ -367,7 +364,7 @@ let rec e_trivial_fail_db only_classes db_list local_db secvars = Proofview.Goal.nf_enter { enter = begin fun gl -> let tacs = e_trivial_resolve db_list local_db secvars only_classes - (project gl) (pf_concl gl) in + (project gl) (EConstr.of_constr (pf_concl gl)) in tclFIRST (List.map (fun (x,_,_,_,_) -> x) tacs) end} in @@ -379,13 +376,13 @@ let rec e_trivial_fail_db only_classes db_list local_db secvars = and e_my_find_search db_list local_db secvars hdc complete only_classes sigma concl = let open Proofview.Notations in - let prods, concl = decompose_prod_assum concl in + let prods, concl = EConstr.decompose_prod_assum sigma concl in let nprods = List.length prods in let freeze = try let cl = Typeclasses.class_info (fst hdc) in if cl.cl_strict then - Evd.evars_of_term concl + Evarutil.undefined_evars_of_term sigma concl else Evar.Set.empty with e when CErrors.noncritical e -> Evar.Set.empty in @@ -394,8 +391,8 @@ and e_my_find_search db_list local_db secvars hdc complete only_classes sigma co (fun db -> let tacs = if Hint_db.use_dn db then (* Using dnet *) - Hint_db.map_eauto secvars hdc concl db - else Hint_db.map_existential secvars hdc concl db + Hint_db.map_eauto sigma secvars hdc concl db + else Hint_db.map_existential sigma secvars hdc concl db in let flags = auto_unif_flags freeze (Hint_db.transparent_state db) in List.map (fun x -> (flags, x)) tacs) @@ -481,7 +478,7 @@ let catchable = function let pr_depth l = prlist_with_sep (fun () -> str ".") int (List.rev l) let is_Prop env sigma concl = - let ty = Retyping.get_type_of env sigma (EConstr.of_constr concl) in + let ty = Retyping.get_type_of env sigma concl in match kind_of_term ty with | Sort (Prop Null) -> true | _ -> false @@ -542,9 +539,10 @@ let make_resolve_hyp env sigma st flags only_classes pri decl = else false in let is_class = iscl env cty in + let cty = EConstr.of_constr cty in let keep = not only_classes || is_class in if keep then - let c = mkVar id in + let c = EConstr.mkVar id in let name = PathHints [VarRef id] in let hints = if is_class then @@ -552,7 +550,7 @@ let make_resolve_hyp env sigma st flags only_classes pri decl = (List.map_append (fun (path,pri, c) -> make_resolves env sigma ~name:(PathHints path) (true,false,Flags.is_verbose()) pri false - (IsConstr (c,Univ.ContextSet.empty))) + (IsConstr (EConstr.of_constr c,Univ.ContextSet.empty))) hints) else [] in @@ -674,17 +672,16 @@ module V85 = struct let needs_backtrack env evd oev concl = if Option.is_empty oev || is_Prop env evd concl then - occur_existential evd (EConstr.of_constr concl) + occur_existential evd concl else true let hints_tac hints sk fk {it = gl,info; sigma = s} = let env = Goal.V82.env s gl in let concl = Goal.V82.concl s gl in - let concl = EConstr.Unsafe.to_constr concl in let tacgl = {it = gl; sigma = s;} in let secvars = secvars_of_hyps (Environ.named_context_of_val (Goal.V82.hyps s gl)) in let poss = e_possible_resolve hints info.hints secvars info.only_classes s concl in - let unique = is_unique env s (EConstr.of_constr concl) in + let unique = is_unique env s concl in let rec aux i foundone = function | (tac, _, extern, name, pp) :: tl -> let derivs = path_derivate info.auto_cut name in @@ -749,7 +746,7 @@ module V85 = struct let fk' = (fun e -> let do_backtrack = - if unique then occur_existential s' (EConstr.of_constr concl) + if unique then occur_existential s' concl else if info.unique then true else if List.is_empty gls' then needs_backtrack env s' info.is_evar concl @@ -770,7 +767,7 @@ module V85 = struct if foundone == None && !typeclasses_debug > 0 then Feedback.msg_debug (pr_depth info.auto_depth ++ str": no match for " ++ - Printer.pr_constr_env (Goal.V82.env s gl) s concl ++ + Printer.pr_constr_env (Goal.V82.env s gl) s (EConstr.Unsafe.to_constr concl) ++ spc () ++ str ", " ++ int (List.length poss) ++ str" possibilities"); match foundone with @@ -793,7 +790,7 @@ module V85 = struct let fk'' = if not info.unique && List.is_empty gls' && not (needs_backtrack (Goal.V82.env s gl) s - info.is_evar (EConstr.Unsafe.to_constr (Goal.V82.concl s gl))) + info.is_evar (Goal.V82.concl s gl)) then fk else fk' in @@ -984,7 +981,7 @@ module Search = struct NOT backtrack. *) let needs_backtrack env evd unique concl = if unique || is_Prop env evd concl then - occur_existential evd (EConstr.of_constr concl) + occur_existential evd concl else true let mark_unresolvables sigma goals = @@ -1004,14 +1001,15 @@ module Search = struct let open Proofview.Notations in let env = Goal.env gl in let concl = Goal.concl gl in + let concl = EConstr.of_constr concl in let sigma = Goal.sigma gl in let s = Sigma.to_evar_map sigma in - let unique = not info.search_dep || is_unique env s (EConstr.of_constr concl) in + let unique = not info.search_dep || is_unique env s concl in let backtrack = needs_backtrack env s unique concl in if !typeclasses_debug > 0 then Feedback.msg_debug (pr_depth info.search_depth ++ str": looking for " ++ - Printer.pr_constr_env (Goal.env gl) s concl ++ + Printer.pr_constr_env (Goal.env gl) s (EConstr.Unsafe.to_constr concl) ++ (if backtrack then str" with backtracking" else str" without backtracking")); let secvars = compute_secvars gl in @@ -1126,7 +1124,7 @@ module Search = struct if !foundone == false && !typeclasses_debug > 0 then Feedback.msg_debug (pr_depth info.search_depth ++ str": no match for " ++ - Printer.pr_constr_env (Goal.env gl) s concl ++ + Printer.pr_constr_env (Goal.env gl) s (EConstr.Unsafe.to_constr concl) ++ spc () ++ str ", " ++ int (List.length poss) ++ str" possibilities"); match e with @@ -1523,8 +1521,9 @@ let is_ground c gl = let autoapply c i gl = let flags = auto_unif_flags Evar.Set.empty (Hints.Hint_db.transparent_state (Hints.searchtable_map i)) in - let cty = pf_unsafe_type_of gl (EConstr.of_constr c) in - let ce = mk_clenv_from gl (EConstr.of_constr c,EConstr.of_constr cty) in + let cty = pf_unsafe_type_of gl c in + let cty = EConstr.of_constr cty in + let ce = mk_clenv_from gl (c,cty) in let tac = { enter = fun gl -> (unify_e_resolve false flags).enter gl ((c,cty,Univ.ContextSet.empty),0,ce) } in Proofview.V82.of_tactic (Proofview.Goal.nf_enter tac) gl diff --git a/tactics/class_tactics.mli b/tactics/class_tactics.mli index 8db264ad9..027b7dcd7 100644 --- a/tactics/class_tactics.mli +++ b/tactics/class_tactics.mli @@ -30,7 +30,7 @@ val not_evar : constr -> unit Proofview.tactic val is_ground : constr -> tactic -val autoapply : constr -> Hints.hint_db_name -> tactic +val autoapply : EConstr.constr -> Hints.hint_db_name -> tactic module Search : sig val eauto_tac : diff --git a/tactics/eauto.ml b/tactics/eauto.ml index 24e4de750..5d42ed2d5 100644 --- a/tactics/eauto.ml +++ b/tactics/eauto.ml @@ -29,7 +29,6 @@ open Proofview.Notations let eauto_unif_flags = auto_flags_of_state full_transparent_state let e_give_exact ?(flags=eauto_unif_flags) c = - let c = EConstr.of_constr c in Proofview.Goal.enter { enter = begin fun gl -> let t1 = Tacmach.New.pf_unsafe_type_of gl c in let t1 = EConstr.of_constr t1 in @@ -40,7 +39,7 @@ let e_give_exact ?(flags=eauto_unif_flags) c = else exact_check c end } -let assumption id = e_give_exact (mkVar id) +let assumption id = e_give_exact (EConstr.mkVar id) let e_assumption = Proofview.Goal.enter { enter = begin fun gl -> @@ -49,7 +48,7 @@ let e_assumption = let registered_e_assumption = Proofview.Goal.enter { enter = begin fun gl -> - Tacticals.New.tclFIRST (List.map (fun id -> e_give_exact (mkVar id)) + Tacticals.New.tclFIRST (List.map (fun id -> e_give_exact (EConstr.mkVar id)) (Tacmach.New.pf_ids_of_hyps gl)) end } @@ -89,15 +88,14 @@ let rec prolog l n gl = let out_term = function | IsConstr (c, _) -> c - | IsGlobRef gr -> fst (Universes.fresh_global_instance (Global.env ()) gr) + | IsGlobRef gr -> EConstr.of_constr (fst (Universes.fresh_global_instance (Global.env ()) gr)) let prolog_tac l n = Proofview.V82.tactic begin fun gl -> let map c = let (c, sigma) = Tactics.run_delayed (pf_env gl) (project gl) c in - let c = EConstr.Unsafe.to_constr c in let c = pf_apply (prepare_hint false (false,true)) gl (sigma, c) in - EConstr.of_constr (out_term c) + out_term c in let l = List.map map l in try (prolog l n gl) @@ -116,7 +114,6 @@ let priority l = List.map snd (List.filter (fun (pr,_) -> Int.equal pr 0) l) let unify_e_resolve poly flags (c,clenv) = Proofview.Goal.nf_enter { enter = begin fun gl -> let clenv', c = connect_hint_clenv poly c clenv gl in - let c = EConstr.of_constr c in Proofview.V82.tactic (fun gls -> let clenv' = clenv_unique_resolver ~flags clenv' gls in @@ -124,13 +121,13 @@ let unify_e_resolve poly flags (c,clenv) = (Proofview.V82.of_tactic (Tactics.Simple.eapply c)) gls) end } -let hintmap_of secvars hdc concl = +let hintmap_of sigma secvars hdc concl = match hdc with | None -> fun db -> Hint_db.map_none ~secvars db | Some hdc -> - if occur_existential Evd.empty (EConstr.of_constr concl) then (** FIXME *) - (fun db -> Hint_db.map_existential ~secvars hdc concl db) - else (fun db -> Hint_db.map_auto ~secvars hdc concl db) + if occur_existential sigma concl then + (fun db -> Hint_db.map_existential sigma ~secvars hdc concl db) + else (fun db -> Hint_db.map_auto sigma ~secvars hdc concl db) (* FIXME: should be (Hint_db.map_eauto hdc concl db) *) let e_exact poly flags (c,clenv) = @@ -152,13 +149,13 @@ let rec e_trivial_fail_db db_list local_db = let tacl = registered_e_assumption :: (Tacticals.New.tclTHEN Tactics.intro next) :: - (List.map fst (e_trivial_resolve (Tacmach.New.project gl) db_list local_db secvars (Tacmach.New.pf_nf_concl gl))) + (List.map fst (e_trivial_resolve (Tacmach.New.project gl) db_list local_db secvars (EConstr.of_constr (Tacmach.New.pf_nf_concl gl)))) in Tacticals.New.tclFIRST (List.map Tacticals.New.tclCOMPLETE tacl) end } -and e_my_find_search db_list local_db secvars hdc concl = - let hint_of_db = hintmap_of secvars hdc concl in +and e_my_find_search sigma db_list local_db secvars hdc concl = + let hint_of_db = hintmap_of sigma secvars hdc concl in let hintl = List.map_append (fun db -> let flags = auto_flags_of_state (Hint_db.transparent_state db) in @@ -188,13 +185,13 @@ and e_my_find_search db_list local_db secvars hdc concl = and e_trivial_resolve sigma db_list local_db secvars gl = let hd = try Some (decompose_app_bound sigma gl) with Bound -> None in - try priority (e_my_find_search db_list local_db secvars hd gl) + try priority (e_my_find_search sigma db_list local_db secvars hd gl) with Not_found -> [] let e_possible_resolve sigma db_list local_db secvars gl = let hd = try Some (decompose_app_bound sigma gl) with Bound -> None in try List.map (fun (b, (tac, pp)) -> (tac, b, pp)) - (e_my_find_search db_list local_db secvars hd gl) + (e_my_find_search sigma db_list local_db secvars hd gl) with Not_found -> [] let find_first_goal gls = @@ -265,7 +262,7 @@ module SearchProblem = struct let g = find_first_goal lg in let hyps = pf_ids_of_hyps g in let secvars = secvars_of_hyps (pf_hyps g) in - let map_assum id = (e_give_exact (mkVar id), (-1), lazy (str "exact" ++ spc () ++ pr_id id)) in + let map_assum id = (e_give_exact (EConstr.mkVar id), (-1), lazy (str "exact" ++ spc () ++ pr_id id)) in let assumption_tacs = let tacs = List.map map_assum hyps in let l = filter_tactics s.tacres tacs in @@ -293,6 +290,7 @@ module SearchProblem = struct let rec_tacs = let l = let concl = Reductionops.nf_evar (project g)(pf_concl g) in + let concl = EConstr.of_constr concl in filter_tactics s.tacres (e_possible_resolve (project g) s.dblist (List.hd s.localdb) secvars concl) in diff --git a/tactics/eauto.mli b/tactics/eauto.mli index 1f69e4ab3..defa34d9c 100644 --- a/tactics/eauto.mli +++ b/tactics/eauto.mli @@ -15,7 +15,7 @@ val e_assumption : unit Proofview.tactic val registered_e_assumption : unit Proofview.tactic -val e_give_exact : ?flags:Unification.unify_flags -> constr -> unit Proofview.tactic +val e_give_exact : ?flags:Unification.unify_flags -> EConstr.constr -> unit Proofview.tactic val prolog_tac : delayed_open_constr list -> int -> unit Proofview.tactic diff --git a/tactics/hints.ml b/tactics/hints.ml index c31e86383..cd5fc79f5 100644 --- a/tactics/hints.ml +++ b/tactics/hints.ml @@ -6,12 +6,16 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +module CVars = Vars + open Pp open Util open CErrors open Names -open Vars open Term +open Evd +open EConstr +open Vars open Environ open Mod_subst open Globnames @@ -21,7 +25,6 @@ open Libnames open Smartlocate open Misctypes open Tactypes -open Evd open Termops open Inductiveops open Typing @@ -46,22 +49,29 @@ type debug = Debug | Info | Off exception Bound let head_constr_bound sigma t = - let t = strip_outer_cast sigma (EConstr.of_constr t) in - let _,ccl = decompose_prod_assum t in - let hd,args = decompose_app ccl in - match kind_of_term hd with - | Const _ | Ind _ | Construct _ | Var _ -> hd - | Proj (p, _) -> mkConst (Projection.constant p) - | _ -> raise Bound + let t = strip_outer_cast sigma t in + let t = EConstr.of_constr t in + let _,ccl = decompose_prod_assum sigma t in + let hd,args = decompose_app sigma ccl in + match EConstr.kind sigma hd with + | Const (c, _) -> ConstRef c + | Ind (i, _) -> IndRef i + | Construct (c, _) -> ConstructRef c + | Var id -> VarRef id + | Proj (p, _) -> ConstRef (Projection.constant p) + | _ -> raise Bound let head_constr sigma c = try head_constr_bound sigma c with Bound -> error "Bound head variable." let decompose_app_bound sigma t = - let t = strip_outer_cast sigma (EConstr.of_constr t) in - let _,ccl = decompose_prod_assum t in - let hd,args = decompose_app_vect sigma (EConstr.of_constr ccl) in - match kind_of_term hd with + let t = strip_outer_cast sigma t in + let t = EConstr.of_constr t in + let _,ccl = decompose_prod_assum sigma t in + let hd,args = decompose_app_vect sigma ccl in + let hd = EConstr.of_constr hd in + let args = Array.map EConstr.of_constr args in + match EConstr.kind sigma hd with | Const (c,u) -> ConstRef c, args | Ind (i,u) -> IndRef i, args | Construct (c,u) -> ConstructRef c, args @@ -245,6 +255,7 @@ let rebuild_dn st se = { se with sentry_bnet = dn' } let lookup_tacs concl st se = + let concl = EConstr.Unsafe.to_constr concl in let l' = Bounded_net.lookup st se.sentry_bnet concl in let sl' = List.stable_sort pri_order_int l' in List.merge pri_order_int se.sentry_nopat sl' @@ -256,10 +267,10 @@ let is_transparent_gr (ids, csts) = function | ConstRef cst -> Cpred.mem cst csts | IndRef _ | ConstructRef _ -> false -let strip_params env c = - match kind_of_term c with +let strip_params env sigma c = + match EConstr.kind sigma c with | App (f, args) -> - (match kind_of_term f with + (match EConstr.kind sigma f with | Const (p,_) -> let cb = lookup_constant p env in (match cb.Declarations.const_proj with @@ -276,11 +287,9 @@ let strip_params env c = let instantiate_hint env sigma p = let mk_clenv (c, cty, ctx) = let sigma = Evd.merge_context_set univ_flexible sigma ctx in - let c = EConstr.of_constr c in - let cty = EConstr.of_constr cty in let cl = mk_clenv_from_env env sigma None (c,cty) in {cl with templval = - { cl.templval with rebus = EConstr.of_constr (strip_params env (EConstr.Unsafe.to_constr cl.templval.rebus)) }; + { cl.templval with rebus = strip_params env sigma cl.templval.rebus }; env = empty_env} in let code = match p.code.obj with @@ -445,11 +454,11 @@ val empty : ?name:hint_db_name -> transparent_state -> bool -> t val find : global_reference -> t -> search_entry val map_none : secvars:Id.Pred.t -> t -> full_hint list val map_all : secvars:Id.Pred.t -> global_reference -> t -> full_hint list -val map_existential : secvars:Id.Pred.t -> +val map_existential : evar_map -> secvars:Id.Pred.t -> (global_reference * constr array) -> constr -> t -> full_hint list -val map_eauto : secvars:Id.Pred.t -> +val map_eauto : evar_map -> secvars:Id.Pred.t -> (global_reference * constr array) -> constr -> t -> full_hint list -val map_auto : secvars:Id.Pred.t -> +val map_auto : evar_map -> secvars:Id.Pred.t -> (global_reference * constr array) -> constr -> t -> full_hint list val add_one : env -> evar_map -> hint_entry -> t -> t val add_list : env -> evar_map -> hint_entry list -> t -> t @@ -505,21 +514,21 @@ struct (** Warn about no longer typable hint? *) None - let match_mode m arg = + let match_mode sigma m arg = match m with - | ModeInput -> not (occur_existential Evd.empty (EConstr.of_constr arg)) (** FIXME *) + | ModeInput -> not (occur_existential sigma arg) | ModeNoHeadEvar -> - Evarutil.(try ignore(head_evar arg); false + Evarutil.(try ignore(head_evar sigma arg); false with NoHeadEvar -> true) | ModeOutput -> true - let matches_mode args mode = + let matches_mode sigma args mode = Array.length mode == Array.length args && - Array.for_all2 match_mode mode args + Array.for_all2 (match_mode sigma) mode args - let matches_modes args modes = + let matches_modes sigma args modes = if List.is_empty modes then true - else List.exists (matches_mode args) modes + else List.exists (matches_mode sigma args) modes let merge_entry secvars db nopat pat = let h = List.sort pri_order_int (List.map snd db.hintdb_nopat) in @@ -535,22 +544,22 @@ struct merge_entry secvars db se.sentry_nopat se.sentry_pat (** Precondition: concl has no existentials *) - let map_auto ~secvars (k,args) concl db = + let map_auto sigma ~secvars (k,args) concl db = let se = find k db in let st = if db.use_dn then (Some db.hintdb_state) else None in let pat = lookup_tacs concl st se in merge_entry secvars db [] pat - let map_existential ~secvars (k,args) concl db = + let map_existential sigma ~secvars (k,args) concl db = let se = find k db in - if matches_modes args se.sentry_mode then + if matches_modes sigma args se.sentry_mode then merge_entry secvars db se.sentry_nopat se.sentry_pat else merge_entry secvars db [] [] (* [c] contains an existential *) - let map_eauto ~secvars (k,args) concl db = + let map_eauto sigma ~secvars (k,args) concl db = let se = find k db in - if matches_modes args se.sentry_mode then + if matches_modes sigma args se.sentry_mode then let st = if db.use_dn then Some db.hintdb_state else None in let pat = lookup_tacs concl st se in merge_entry secvars db [] pat @@ -718,8 +727,8 @@ let _ = Summary.declare_summary "search" (* Auxiliary functions to prepare AUTOHINT objects *) (**************************************************************************) -let rec nb_hyp c = match kind_of_term c with - | Prod(_,_,c2) -> if noccurn 1 c2 then 1+(nb_hyp c2) else nb_hyp c2 +let rec nb_hyp sigma c = match EConstr.kind sigma c with + | Prod(_,_,c2) -> if noccurn sigma 1 c2 then 1+(nb_hyp sigma c2) else nb_hyp sigma c2 | _ -> 0 (* adding and removing tactics in the search table *) @@ -736,19 +745,20 @@ let secvars_of_idset s = Id.Pred.add id p else p) s Id.Pred.empty -let secvars_of_constr env c = - secvars_of_idset (Environ.global_vars_set env c) +let secvars_of_constr env sigma c = + secvars_of_idset (Termops.global_vars_set env sigma c) let secvars_of_global env gr = secvars_of_idset (vars_of_global_reference env gr) let make_exact_entry env sigma pri poly ?(name=PathAny) (c, cty, ctx) = - let secvars = secvars_of_constr env c in - let cty = strip_outer_cast sigma (EConstr.of_constr cty) in - match kind_of_term cty with + let secvars = secvars_of_constr env sigma c in + let cty = strip_outer_cast sigma cty in + let cty = EConstr.of_constr cty in + match EConstr.kind sigma cty with | Prod _ -> failwith "make_exact_entry" | _ -> - let pat = Patternops.pattern_of_constr env sigma (EConstr.of_constr cty) in + let pat = Patternops.pattern_of_constr env sigma cty in let hd = try head_pattern_bound pat with BoundPattern -> failwith "make_exact_entry" @@ -763,21 +773,21 @@ let make_exact_entry env sigma pri poly ?(name=PathAny) (c, cty, ctx) = code = with_uid (Give_exact (c, cty, ctx)); }) let make_apply_entry env sigma (eapply,hnf,verbose) pri poly ?(name=PathAny) (c, cty, ctx) = - let cty = if hnf then hnf_constr env sigma (EConstr.of_constr cty) else cty in - match kind_of_term cty with + let cty = if hnf then EConstr.of_constr (hnf_constr env sigma cty) else cty in + match EConstr.kind sigma cty with | Prod _ -> let sigma' = Evd.merge_context_set univ_flexible sigma ctx in - let ce = mk_clenv_from_env env sigma' None (EConstr.of_constr c,EConstr.of_constr cty) in + let ce = mk_clenv_from_env env sigma' None (c,cty) in let c' = clenv_type (* ~reduce:false *) ce in let pat = Patternops.pattern_of_constr env ce.evd c' in let hd = try head_pattern_bound pat with BoundPattern -> failwith "make_apply_entry" in let nmiss = List.length (clenv_missing ce) in - let secvars = secvars_of_constr env c in + let secvars = secvars_of_constr env sigma c in if Int.equal nmiss 0 then (Some hd, - { pri = (match pri with None -> nb_hyp cty | Some p -> p); + { pri = (match pri with None -> nb_hyp sigma' cty | Some p -> p); poly = poly; pat = Some pat; name = name; @@ -787,10 +797,10 @@ let make_apply_entry env sigma (eapply,hnf,verbose) pri poly ?(name=PathAny) (c, else begin if not eapply then failwith "make_apply_entry"; if verbose then - Feedback.msg_info (str "the hint: eapply " ++ pr_lconstr c ++ + Feedback.msg_info (str "the hint: eapply " ++ pr_lconstr (EConstr.Unsafe.to_constr c) ++ str " will only be used by eauto"); (Some hd, - { pri = (match pri with None -> nb_hyp cty + nmiss | Some p -> p); + { pri = (match pri with None -> nb_hyp sigma' cty + nmiss | Some p -> p); poly = poly; pat = Some pat; name = name; @@ -808,7 +818,7 @@ let pr_hint_term env sigma ctx = function | IsGlobRef gr -> pr_global gr | IsConstr (c, ctx) -> let sigma = Evd.merge_context_set Evd.univ_flexible sigma ctx in - pr_constr_env env sigma c + pr_constr_env env sigma (EConstr.Unsafe.to_constr c) (** We need an object to record the side-effect of registering global universes associated with a hint. *) @@ -834,7 +844,8 @@ let fresh_global_or_constr env sigma poly cr = let isgr, (c, ctx) = match cr with | IsGlobRef gr -> - true, Universes.fresh_global_instance env gr + let (c, ctx) = Universes.fresh_global_instance env gr in + true, (EConstr.of_constr c, ctx) | IsConstr (c, ctx) -> false, (c, ctx) in if poly then (c, ctx) @@ -848,7 +859,8 @@ let fresh_global_or_constr env sigma poly cr = let make_resolves env sigma flags pri poly ?name cr = let c, ctx = fresh_global_or_constr env sigma poly cr in - let cty = Retyping.get_type_of env sigma (EConstr.of_constr c) in + let cty = Retyping.get_type_of env sigma c in + let cty = EConstr.of_constr cty in let try_apply f = try Some (f (c, cty, ctx)) with Failure _ -> None in let ents = List.map_filter try_apply @@ -857,7 +869,7 @@ let make_resolves env sigma flags pri poly ?name cr = in if List.is_empty ents then user_err ~hdr:"Hint" - (pr_lconstr c ++ spc() ++ + (pr_lconstr (EConstr.Unsafe.to_constr c) ++ spc() ++ (if pi1 flags then str"cannot be used as a hint." else str "can be used as a hint only for eauto.")); ents @@ -869,7 +881,7 @@ let make_resolve_hyp env sigma decl = try [make_apply_entry env sigma (true, true, false) None false ~name:(PathHints [VarRef hname]) - (c, NamedDecl.get_type decl, Univ.ContextSet.empty)] + (c, EConstr.of_constr (NamedDecl.get_type decl), Univ.ContextSet.empty)] with | Failure _ -> [] | e when Logic.catchable_exception e -> anomaly (Pp.str "make_resolve_hyp") @@ -899,6 +911,7 @@ let make_extern pri pat tacast = code = with_uid (Extern tacast) }) let make_mode ref m = + let open Term in let ty = Global.type_of_global_unsafe ref in let ctx, t = decompose_prod ty in let n = List.length ctx in @@ -912,15 +925,16 @@ let make_mode ref m = let make_trivial env sigma poly ?(name=PathAny) r = let c,ctx = fresh_global_or_constr env sigma poly r in let sigma = Evd.merge_context_set univ_flexible sigma ctx in - let t = hnf_constr env sigma (EConstr.of_constr (unsafe_type_of env sigma (EConstr.of_constr c))) in - let hd = head_of_constr_reference sigma (EConstr.of_constr (head_constr sigma t)) in - let ce = mk_clenv_from_env env sigma None (EConstr.of_constr c,EConstr.of_constr t) in + let t = hnf_constr env sigma (EConstr.of_constr (unsafe_type_of env sigma c)) in + let t = EConstr.of_constr t in + let hd = head_constr sigma t in + let ce = mk_clenv_from_env env sigma None (c,t) in (Some hd, { pri=1; poly = poly; pat = Some (Patternops.pattern_of_constr env ce.evd (clenv_type ce)); name = name; db = None; - secvars = secvars_of_constr env c; + secvars = secvars_of_constr env sigma c; code= with_uid (Res_pf_THEN_trivial_fail(c,t,ctx)) }) @@ -1014,14 +1028,16 @@ let cache_autohint (kn, obj) = let subst_autohint (subst, obj) = let subst_key gr = let (lab'', elab') = subst_global subst gr in + let elab' = EConstr.of_constr elab' in let gr' = - (try head_of_constr_reference Evd.empty (EConstr.of_constr (head_constr_bound Evd.empty (** FIXME *) elab')) + (try head_constr_bound Evd.empty elab' with Bound -> lab'') in if gr' == gr then gr else gr' in let subst_hint (k,data as hint) = let k' = Option.smartmap subst_key k in let pat' = Option.smartmap (subst_pattern subst) data.pat in + let subst_mps subst c = EConstr.of_constr (subst_mps subst (EConstr.Unsafe.to_constr c)) in let code' = match data.code.obj with | Res_pf (c,t,ctx) -> let c' = subst_mps subst c in @@ -1191,31 +1207,33 @@ let prepare_hint check (poly,local) env init (sigma,c) = It is actually a bit stupid to generalize over evars since the first thing make_resolves will do is to re-instantiate the products *) let sigma, subst = Evd.nf_univ_variables sigma in - let c = Vars.subst_univs_constr subst (Evarutil.nf_evar sigma c) in - let c = drop_extra_implicit_args sigma (EConstr.of_constr c) in - let vars = ref (collect_vars sigma (EConstr.of_constr c)) in + let c = Evarutil.nf_evar sigma (EConstr.Unsafe.to_constr c) in + let c = CVars.subst_univs_constr subst c in + let c = EConstr.of_constr c in + let c = drop_extra_implicit_args sigma c in + let vars = ref (collect_vars sigma c) in let subst = ref [] in - let rec find_next_evar c = match kind_of_term c with + let rec find_next_evar c = match EConstr.kind sigma c with | Evar (evk,args as ev) -> (* We skip the test whether args is the identity or not *) - let t = Evarutil.nf_evar sigma (existential_type sigma ev) in - let t = List.fold_right (fun (e,id) c -> replace_term sigma (EConstr.of_constr e) (EConstr.of_constr id) (EConstr.of_constr c)) !subst t in - if not (closed0 c) then + let t = existential_type sigma ev in + let t = List.fold_right (fun (e,id) c -> EConstr.of_constr (replace_term sigma e id c)) !subst t in + if not (closed0 sigma c) then error "Hints with holes dependent on a bound variable not supported."; - if occur_existential sigma (EConstr.of_constr t) then + if occur_existential sigma t then (* Not clever enough to construct dependency graph of evars *) error "Not clever enough to deal with evars dependent in other evars."; raise (Found (c,t)) - | _ -> Constr.iter find_next_evar c in + | _ -> EConstr.iter sigma find_next_evar c in let rec iter c = try find_next_evar c; c with Found (evar,t) -> let id = next_ident_away_from default_prepare_hint_ident (fun id -> Id.Set.mem id !vars) in vars := Id.Set.add id !vars; subst := (evar,mkVar id)::!subst; - mkNamedLambda id t (iter (replace_term sigma (EConstr.of_constr evar) (EConstr.mkVar id) (EConstr.of_constr c))) in + mkNamedLambda id t (iter (EConstr.of_constr (replace_term sigma evar (mkVar id) c))) in let c' = iter c in - if check then Pretyping.check_evars (Global.env()) Evd.empty sigma (EConstr.of_constr c'); + if check then Pretyping.check_evars (Global.env()) Evd.empty sigma c'; let diff = Univ.ContextSet.diff (Evd.universe_context_set sigma) (Evd.universe_context_set init) in if poly then IsConstr (c', diff) else if local then IsConstr (c', diff) @@ -1228,6 +1246,7 @@ let interp_hints poly = let sigma = Evd.from_env env in let f poly c = let evd,c = Constrintern.interp_open_constr env sigma c in + let c = EConstr.of_constr c in prepare_hint true (poly,false) (Global.env()) Evd.empty (evd,c) in let fref r = let gr = global_with_alias r in @@ -1293,7 +1312,7 @@ let add_hints local dbnames0 h = let expand_constructor_hints env sigma lems = List.map_append (fun (evd,lem) -> - match kind_of_term lem with + match EConstr.kind sigma lem with | Ind (ind,u) -> List.init (nconstructors ind) (fun i -> @@ -1320,7 +1339,7 @@ let make_local_hint_db env sigma ts eapply lems = let map c = let sigma = Sigma.Unsafe.of_evar_map sigma in let Sigma (c, sigma, _) = c.delayed env sigma in - (Sigma.to_evar_map sigma, EConstr.Unsafe.to_constr c) + (Sigma.to_evar_map sigma, c) in let lems = List.map map lems in let sign = Environ.named_context env in @@ -1348,7 +1367,7 @@ let make_db_list dbnames = (* Functions for printing the hints *) (**************************************************************************) -let pr_hint_elt (c, _, _) = pr_constr c +let pr_hint_elt (c, _, _) = pr_constr (EConstr.Unsafe.to_constr c) let pr_hint h = match h.obj with | Res_pf (c, _) -> (str"simple apply " ++ pr_hint_elt c) @@ -1402,9 +1421,9 @@ let pr_hint_term sigma cl = let valid_dbs = let fn = try let hdc = decompose_app_bound sigma cl in - if occur_existential sigma (EConstr.of_constr cl) then - Hint_db.map_existential ~secvars:Id.Pred.full hdc cl - else Hint_db.map_auto ~secvars:Id.Pred.full hdc cl + if occur_existential sigma cl then + Hint_db.map_existential sigma ~secvars:Id.Pred.full hdc cl + else Hint_db.map_auto sigma ~secvars:Id.Pred.full hdc cl with Bound -> Hint_db.map_none ~secvars:Id.Pred.full in let fn db = List.map (fun x -> 0, x) (fn db) in @@ -1425,7 +1444,7 @@ let pr_applicable_hint () = match glss.Evd.it with | [] -> CErrors.error "No focused goal." | g::_ -> - pr_hint_term glss.Evd.sigma (EConstr.Unsafe.to_constr (Goal.V82.concl glss.Evd.sigma g)) + pr_hint_term glss.Evd.sigma (Goal.V82.concl glss.Evd.sigma g) let pp_hint_mode = function | ModeInput -> str"+" diff --git a/tactics/hints.mli b/tactics/hints.mli index c0eb2c3b8..344827e03 100644 --- a/tactics/hints.mli +++ b/tactics/hints.mli @@ -10,6 +10,7 @@ open Pp open Util open Names open Term +open EConstr open Environ open Globnames open Decl_kinds @@ -99,16 +100,16 @@ module Hint_db : (** All hints associated to the reference, respecting modes if evars appear in the arguments, _not_ using the discrimination net. *) - val map_existential : secvars:Id.Pred.t -> + val map_existential : evar_map -> secvars:Id.Pred.t -> (global_reference * constr array) -> constr -> t -> full_hint list (** All hints associated to the reference, respecting modes if evars appear in the arguments and using the discrimination net. *) - val map_eauto : secvars:Id.Pred.t -> (global_reference * constr array) -> constr -> t -> full_hint list + val map_eauto : evar_map -> secvars:Id.Pred.t -> (global_reference * constr array) -> constr -> t -> full_hint list (** All hints associated to the reference, respecting modes if evars appear in the arguments. *) - val map_auto : secvars:Id.Pred.t -> + val map_auto : evar_map -> secvars:Id.Pred.t -> (global_reference * constr array) -> constr -> t -> full_hint list val add_one : env -> evar_map -> hint_entry -> t -> t @@ -170,7 +171,7 @@ val add_hints : locality_flag -> hint_db_name list -> hints_entry -> unit val prepare_hint : bool (* Check no remaining evars *) -> (bool * bool) (* polymorphic or monomorphic, local or global *) -> - env -> evar_map -> open_constr -> hint_term + env -> evar_map -> evar_map * constr -> hint_term (** [make_exact_entry pri (c, ctyp, ctx, secvars)]. [c] is the term given as an exact proof to solve the goal; diff --git a/toplevel/classes.ml b/toplevel/classes.ml index 7c4ad6225..acfbce357 100644 --- a/toplevel/classes.ml +++ b/toplevel/classes.ml @@ -49,7 +49,7 @@ let set_typeclass_transparency c local b = let _ = Hook.set Typeclasses.add_instance_hint_hook (fun inst path local pri poly -> - let inst' = match inst with IsConstr c -> Hints.IsConstr (c, Univ.ContextSet.empty) + let inst' = match inst with IsConstr c -> Hints.IsConstr (EConstr.of_constr c, Univ.ContextSet.empty) | IsGlobal gr -> Hints.IsGlobRef gr in Flags.silently (fun () -> -- cgit v1.2.3 From 372b092aea6d077fe0a82eac74d742da8998c7ad Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sun, 20 Nov 2016 00:27:08 +0100 Subject: Auto API using EConstr. --- tactics/auto.ml | 9 +++++---- tactics/auto.mli | 5 +++-- 2 files changed, 8 insertions(+), 6 deletions(-) diff --git a/tactics/auto.ml b/tactics/auto.ml index 2423ea878..c34f9dd92 100644 --- a/tactics/auto.ml +++ b/tactics/auto.ml @@ -6,14 +6,15 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -(* -*) +module CVars = Vars + open Pp open Util open CErrors open Names open Vars open Termops +open EConstr open Environ open Tacmach open Genredexpr @@ -83,8 +84,8 @@ let connect_hint_clenv poly (c, _, ctx) clenv gl = if poly then (** Refresh the instance of the hint *) let (subst, ctx) = Universes.fresh_universe_context_set_instance ctx in - let map c = Vars.subst_univs_level_constr subst c in - let emap c = EConstr.Vars.subst_univs_level_constr subst c in + let map c = CVars.subst_univs_level_constr subst c in + let emap c = Vars.subst_univs_level_constr subst c in let evd = Evd.merge_context_set Evd.univ_flexible evd ctx in (** Only metas are mentioning the old universes. *) let clenv = { diff --git a/tactics/auto.mli b/tactics/auto.mli index 403be9e1c..b0ecb4660 100644 --- a/tactics/auto.mli +++ b/tactics/auto.mli @@ -10,6 +10,7 @@ open Names open Term +open EConstr open Clenv open Pattern open Decl_kinds @@ -23,7 +24,7 @@ val default_search_depth : int ref val auto_flags_of_state : transparent_state -> Unification.unify_flags val connect_hint_clenv : polymorphic -> raw_hint -> clausenv -> - ('a, 'r) Proofview.Goal.t -> clausenv * EConstr.constr + ('a, 'r) Proofview.Goal.t -> clausenv * constr (** Try unification with the precompiled clause, then use registered Apply *) val unify_resolve : polymorphic -> Unification.unify_flags -> (raw_hint * clausenv) -> unit Proofview.tactic @@ -33,7 +34,7 @@ val unify_resolve : polymorphic -> Unification.unify_flags -> (raw_hint * clause [Pattern.somatches], then replace [?1] [?2] metavars in tacast by the right values to build a tactic *) -val conclPattern : EConstr.constr -> constr_pattern option -> Genarg.glob_generic_argument -> unit Proofview.tactic +val conclPattern : constr -> constr_pattern option -> Genarg.glob_generic_argument -> unit Proofview.tactic (** The Auto tactic *) -- cgit v1.2.3 From e71f6d24465ea7812e9b893ed8e0deb2a44ce4f8 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sun, 20 Nov 2016 00:28:38 +0100 Subject: Eauto API using EConstr. --- tactics/eauto.ml | 31 +++++++++++++++++-------------- tactics/eauto.mli | 3 ++- 2 files changed, 19 insertions(+), 15 deletions(-) diff --git a/tactics/eauto.ml b/tactics/eauto.ml index 5d42ed2d5..986d1a624 100644 --- a/tactics/eauto.ml +++ b/tactics/eauto.ml @@ -13,6 +13,7 @@ open Names open Nameops open Term open Termops +open EConstr open Proof_type open Tacticals open Tacmach @@ -33,13 +34,14 @@ let e_give_exact ?(flags=eauto_unif_flags) c = let t1 = Tacmach.New.pf_unsafe_type_of gl c in let t1 = EConstr.of_constr t1 in let t2 = Tacmach.New.pf_concl (Proofview.Goal.assume gl) in + let t2 = EConstr.of_constr t2 in let sigma = Tacmach.New.project gl in - if occur_existential sigma t1 || occur_existential sigma (EConstr.of_constr t2) then + if occur_existential sigma t1 || occur_existential sigma t2 then Tacticals.New.tclTHEN (Clenvtac.unify ~flags t1) (exact_no_check c) else exact_check c end } -let assumption id = e_give_exact (EConstr.mkVar id) +let assumption id = e_give_exact (mkVar id) let e_assumption = Proofview.Goal.enter { enter = begin fun gl -> @@ -48,7 +50,7 @@ let e_assumption = let registered_e_assumption = Proofview.Goal.enter { enter = begin fun gl -> - Tacticals.New.tclFIRST (List.map (fun id -> e_give_exact (EConstr.mkVar id)) + Tacticals.New.tclFIRST (List.map (fun id -> e_give_exact (mkVar id)) (Tacmach.New.pf_ids_of_hyps gl)) end } @@ -77,7 +79,7 @@ let apply_tac_list tac glls = let one_step l gl = [Proofview.V82.of_tactic Tactics.intro] - @ (List.map (fun c -> Proofview.V82.of_tactic (Tactics.Simple.eapply c)) (List.map EConstr.mkVar (pf_ids_of_hyps gl))) + @ (List.map (fun c -> Proofview.V82.of_tactic (Tactics.Simple.eapply c)) (List.map mkVar (pf_ids_of_hyps gl))) @ (List.map (fun c -> Proofview.V82.of_tactic (Tactics.Simple.eapply c)) l) @ (List.map (fun c -> Proofview.V82.of_tactic (assumption c)) (pf_ids_of_hyps gl)) @@ -262,7 +264,7 @@ module SearchProblem = struct let g = find_first_goal lg in let hyps = pf_ids_of_hyps g in let secvars = secvars_of_hyps (pf_hyps g) in - let map_assum id = (e_give_exact (EConstr.mkVar id), (-1), lazy (str "exact" ++ spc () ++ pr_id id)) in + let map_assum id = (e_give_exact (mkVar id), (-1), lazy (str "exact" ++ spc () ++ pr_id id)) in let assumption_tacs = let tacs = List.map map_assum hyps in let l = filter_tactics s.tacres tacs in @@ -469,18 +471,18 @@ let autounfold_tac db cls = in autounfold dbs cls -let unfold_head env (ids, csts) c = +let unfold_head env sigma (ids, csts) c = let rec aux c = - match kind_of_term c with + match EConstr.kind sigma c with | Var id when Id.Set.mem id ids -> (match Environ.named_body id env with - | Some b -> true, b + | Some b -> true, EConstr.of_constr b | None -> false, c) | Const (cst,u as c) when Cset.mem cst csts -> - true, Environ.constant_value_in env c + true, EConstr.of_constr (Environ.constant_value_in env c) | App (f, args) -> (match aux f with - | true, f' -> true, Reductionops.whd_betaiota Evd.empty (EConstr.of_constr (mkApp (f', args))) + | true, f' -> true, EConstr.of_constr (Reductionops.whd_betaiota sigma (mkApp (f', args))) | false, _ -> let done_, args' = Array.fold_left_i (fun i (done_, acc) arg -> @@ -494,7 +496,7 @@ let unfold_head env (ids, csts) c = else false, c) | _ -> let done_ = ref false in - let c' = map_constr (fun c -> + let c' = EConstr.map sigma (fun c -> if !done_ then c else let x, c' = aux c in done_ := x; c') c @@ -504,7 +506,9 @@ let unfold_head env (ids, csts) c = let autounfold_one db cl = Proofview.Goal.nf_enter { enter = begin fun gl -> let env = Proofview.Goal.env gl in + let sigma = Tacmach.New.project gl in let concl = Proofview.Goal.concl gl in + let concl = EConstr.of_constr concl in let st = List.fold_left (fun (i,c) dbname -> let db = try searchtable_map dbname @@ -513,10 +517,9 @@ let autounfold_one db cl = let (ids, csts) = Hint_db.unfolds db in (Id.Set.union ids i, Cset.union csts c)) (Id.Set.empty, Cset.empty) db in - let did, c' = unfold_head env st - (match cl with Some (id, _) -> Tacmach.New.pf_get_hyp_typ id gl | None -> concl) + let did, c' = unfold_head env sigma st + (match cl with Some (id, _) -> EConstr.of_constr (Tacmach.New.pf_get_hyp_typ id gl) | None -> concl) in - let c' = EConstr.of_constr c' in if did then match cl with | Some hyp -> change_in_hyp None (make_change_arg c') hyp diff --git a/tactics/eauto.mli b/tactics/eauto.mli index defa34d9c..e2006ac1e 100644 --- a/tactics/eauto.mli +++ b/tactics/eauto.mli @@ -7,6 +7,7 @@ (************************************************************************) open Term +open EConstr open Proof_type open Hints open Tactypes @@ -15,7 +16,7 @@ val e_assumption : unit Proofview.tactic val registered_e_assumption : unit Proofview.tactic -val e_give_exact : ?flags:Unification.unify_flags -> EConstr.constr -> unit Proofview.tactic +val e_give_exact : ?flags:Unification.unify_flags -> constr -> unit Proofview.tactic val prolog_tac : delayed_open_constr list -> int -> unit Proofview.tactic -- cgit v1.2.3 From 67507df457be05ee5b651a423031a8cd584934ef Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sun, 20 Nov 2016 00:35:43 +0100 Subject: Class_tactics API using EConstr. --- ltac/g_class.ml4 | 6 +++--- ltac/rewrite.ml | 2 ++ pretyping/typeclasses.mli | 4 ++-- tactics/class_tactics.ml | 38 +++++++++++++++++++------------------- tactics/class_tactics.mli | 5 +++-- tactics/tactics.ml | 1 - 6 files changed, 29 insertions(+), 27 deletions(-) diff --git a/ltac/g_class.ml4 b/ltac/g_class.ml4 index ea9a2b6e1..45ee86c4e 100644 --- a/ltac/g_class.ml4 +++ b/ltac/g_class.ml4 @@ -62,15 +62,15 @@ TACTIC EXTEND typeclasses_eauto END TACTIC EXTEND head_of_constr - [ "head_of_constr" ident(h) constr(c) ] -> [ head_of_constr h c ] + [ "head_of_constr" ident(h) constr(c) ] -> [ head_of_constr h (EConstr.of_constr c) ] END TACTIC EXTEND not_evar - [ "not_evar" constr(ty) ] -> [ not_evar ty ] + [ "not_evar" constr(ty) ] -> [ not_evar (EConstr.of_constr ty) ] END TACTIC EXTEND is_ground - [ "is_ground" constr(ty) ] -> [ Proofview.V82.tactic (is_ground ty) ] + [ "is_ground" constr(ty) ] -> [ Proofview.V82.tactic (is_ground (EConstr.of_constr ty)) ] END TACTIC EXTEND autoapply diff --git a/ltac/rewrite.ml b/ltac/rewrite.ml index f2ffb0413..a2b6cb97c 100644 --- a/ltac/rewrite.ml +++ b/ltac/rewrite.ml @@ -132,6 +132,7 @@ let find_class_proof proof_type proof_method env evars carrier relation = try let evars, goal = app_poly_check env evars proof_type [| carrier ; relation |] in let evars', c = Typeclasses.resolve_one_typeclass env (goalevars evars) (EConstr.of_constr goal) in + let c = EConstr.Unsafe.to_constr c in if extends_undefined (goalevars evars) evars' then raise Not_found else app_poly_check env (evars',cstrevars evars) proof_method [| carrier; relation; c |] with e when Logic.catchable_exception e -> raise Not_found @@ -1951,6 +1952,7 @@ let default_morphism sign m = in let evars, morph = app_poly_check env evars PropGlobal.proper_type [| t; sign; m |] in let evars, mor = resolve_one_typeclass env (goalevars evars) (EConstr.of_constr morph) in + let mor = EConstr.Unsafe.to_constr mor in mor, proper_projection mor morph let add_setoid global binders a aeq t n = diff --git a/pretyping/typeclasses.mli b/pretyping/typeclasses.mli index ec36c57e0..e95aba695 100644 --- a/pretyping/typeclasses.mli +++ b/pretyping/typeclasses.mli @@ -103,7 +103,7 @@ val is_class_type : evar_map -> EConstr.types -> bool val resolve_typeclasses : ?fast_path:bool -> ?filter:evar_filter -> ?unique:bool -> ?split:bool -> ?fail:bool -> env -> evar_map -> evar_map -val resolve_one_typeclass : ?unique:bool -> env -> evar_map -> EConstr.types -> open_constr +val resolve_one_typeclass : ?unique:bool -> env -> evar_map -> EConstr.types -> evar_map * EConstr.constr val set_typeclass_transparency_hook : (evaluable_global_reference -> bool (*local?*) -> bool -> unit) Hook.t val set_typeclass_transparency : evaluable_global_reference -> bool -> bool -> unit @@ -120,7 +120,7 @@ val add_instance_hint : global_reference_or_constr -> global_reference list -> val remove_instance_hint : global_reference -> unit val solve_all_instances_hook : (env -> evar_map -> evar_filter -> bool -> bool -> bool -> evar_map) Hook.t -val solve_one_instance_hook : (env -> evar_map -> EConstr.types -> bool -> open_constr) Hook.t +val solve_one_instance_hook : (env -> evar_map -> EConstr.types -> bool -> evar_map * EConstr.constr) Hook.t val declare_instance : int option -> bool -> global_reference -> unit diff --git a/tactics/class_tactics.ml b/tactics/class_tactics.ml index 8ecdd01f2..0ca661557 100644 --- a/tactics/class_tactics.ml +++ b/tactics/class_tactics.ml @@ -18,6 +18,7 @@ open Util open Names open Term open Termops +open EConstr open Reduction open Proof_type open Tacticals @@ -217,7 +218,6 @@ let auto_unif_flags freeze st = } let e_give_exact flags poly (c,clenv) gl = - let open EConstr in let (c, _, _) = c in let c, gl = if poly then @@ -245,7 +245,6 @@ let unify_resolve poly flags = { enter = begin fun gls (c,_,clenv) -> (** Application of a lemma using [refine] instead of the old [w_unify] *) let unify_resolve_refine poly flags = - let open EConstr in let open Clenv in { enter = begin fun gls ((c, t, ctx),n,clenv) -> let env = Proofview.Goal.env gls in @@ -479,7 +478,8 @@ let pr_depth l = prlist_with_sep (fun () -> str ".") int (List.rev l) let is_Prop env sigma concl = let ty = Retyping.get_type_of env sigma concl in - match kind_of_term ty with + let ty = EConstr.of_constr ty in + match EConstr.kind sigma ty with | Sort (Prop Null) -> true | _ -> false @@ -527,22 +527,23 @@ let evars_to_goals p evm = let make_resolve_hyp env sigma st flags only_classes pri decl = let id = NamedDecl.get_id decl in let cty = Evarutil.nf_evar sigma (NamedDecl.get_type decl) in + let cty = EConstr.of_constr cty in let rec iscl env ty = - let ctx, ar = decompose_prod_assum ty in - match kind_of_term (fst (decompose_app ar)) with + let ctx, ar = decompose_prod_assum sigma ty in + match EConstr.kind sigma (fst (decompose_app sigma ar)) with | Const (c,_) -> is_class (ConstRef c) | Ind (i,_) -> is_class (IndRef i) | _ -> let env' = Environ.push_rel_context ctx env in - let ty' = whd_all env' ar in - if not (Term.eq_constr ty' ar) then iscl env' ty' + let ty' = Reductionops.whd_all env' sigma ar in + let ty' = EConstr.of_constr ty' in + if not (EConstr.eq_constr sigma ty' ar) then iscl env' ty' else false in let is_class = iscl env cty in - let cty = EConstr.of_constr cty in let keep = not only_classes || is_class in if keep then - let c = EConstr.mkVar id in + let c = mkVar id in let name = PathHints [VarRef id] in let hints = if is_class then @@ -1466,6 +1467,7 @@ let resolve_one_typeclass env ?(sigma=Evd.empty) gl unique = let nc, gl, subst, _, _ = Evarutil.push_rel_context_to_named_context env gl in let (gl,t,sigma) = Goal.V82.mk_goal sigma nc gl Store.empty in + let (ev, _) = destEvar sigma t in let gls = { it = gl ; sigma = sigma; } in let hints = searchtable_map typeclasses_db in let st = Hint_db.transparent_state hints in @@ -1480,11 +1482,9 @@ let resolve_one_typeclass env ?(sigma=Evd.empty) gl unique = with Refiner.FailError _ -> raise Not_found in let evd = sig_sig gls' in - let t = EConstr.Unsafe.to_constr t in - let t' = let (ev, inst) = destEvar t in - mkEvar (ev, Array.map_of_list EConstr.Unsafe.to_constr subst) - in - let term = Evarutil.nf_evar evd t' in + let t' = mkEvar (ev, Array.of_list subst) in + let term = Evarutil.nf_evar evd (EConstr.Unsafe.to_constr t') in + let term = EConstr.of_constr term in evd, term let _ = @@ -1495,8 +1495,9 @@ let _ = Used in the partial application tactic. *) let rec head_of_constr sigma t = - let t = strip_outer_cast sigma (EConstr.of_constr (collapse_appl sigma (EConstr.of_constr t))) in - match kind_of_term t with + let t = strip_outer_cast sigma (EConstr.of_constr (collapse_appl sigma t)) in + let t = EConstr.of_constr t in + match EConstr.kind sigma t with | Prod (_,_,c2) -> head_of_constr sigma c2 | LetIn (_,_,_,c2) -> head_of_constr sigma c2 | App (f,args) -> head_of_constr sigma f @@ -1505,17 +1506,16 @@ let rec head_of_constr sigma t = let head_of_constr h c = Proofview.tclEVARMAP >>= fun sigma -> let c = head_of_constr sigma c in - let c = EConstr.of_constr c in letin_tac None (Name h) c None Locusops.allHyps let not_evar c = Proofview.tclEVARMAP >>= fun sigma -> - match Evarutil.kind_of_term_upto sigma c with + match EConstr.kind sigma c with | Evar _ -> Tacticals.New.tclFAIL 0 (str"Evar") | _ -> Proofview.tclUNIT () let is_ground c gl = - if Evarutil.is_ground_term (project gl) (EConstr.of_constr c) then tclIDTAC gl + if Evarutil.is_ground_term (project gl) c then tclIDTAC gl else tclFAIL 0 (str"Not ground") gl let autoapply c i gl = diff --git a/tactics/class_tactics.mli b/tactics/class_tactics.mli index 027b7dcd7..171b5c4ea 100644 --- a/tactics/class_tactics.mli +++ b/tactics/class_tactics.mli @@ -10,6 +10,7 @@ open Names open Constr +open EConstr open Tacmach val catchable : exn -> bool @@ -24,13 +25,13 @@ val typeclasses_eauto : ?only_classes:bool -> ?st:transparent_state -> depth:(Int.t option) -> Hints.hint_db_name list -> unit Proofview.tactic -val head_of_constr : Id.t -> Term.constr -> unit Proofview.tactic +val head_of_constr : Id.t -> constr -> unit Proofview.tactic val not_evar : constr -> unit Proofview.tactic val is_ground : constr -> tactic -val autoapply : EConstr.constr -> Hints.hint_db_name -> tactic +val autoapply : constr -> Hints.hint_db_name -> tactic module Search : sig val eauto_tac : diff --git a/tactics/tactics.ml b/tactics/tactics.ml index 59ffd8b62..c025ca9b5 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -1734,7 +1734,6 @@ let solve_remaining_apply_goals = let concl = EConstr.of_constr concl in if Typeclasses.is_class_type evd concl then let evd', c' = Typeclasses.resolve_one_typeclass env evd concl in - let c' = EConstr.of_constr c' in let tac = Refine.refine ~unsafe:true { run = fun h -> Sigma.here c' h } in Sigma.Unsafe.of_pair (tac, evd') else Sigma.here (Proofview.tclUNIT ()) sigma -- cgit v1.2.3 From c72bf7330bb32970616be4dddc7571f3b91c1562 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sun, 20 Nov 2016 01:08:53 +0100 Subject: Eqdecide API using EConstr. --- ltac/g_eqdecide.ml4 | 2 +- proofs/tacmach.ml | 2 ++ proofs/tacmach.mli | 4 ++-- tactics/eqdecide.ml | 35 ++++++++++++++++++----------------- tactics/eqdecide.mli | 2 +- tactics/hipattern.ml | 5 ++++- tactics/hipattern.mli | 2 +- 7 files changed, 29 insertions(+), 23 deletions(-) diff --git a/ltac/g_eqdecide.ml4 b/ltac/g_eqdecide.ml4 index 905653281..6a49ea830 100644 --- a/ltac/g_eqdecide.ml4 +++ b/ltac/g_eqdecide.ml4 @@ -23,5 +23,5 @@ TACTIC EXTEND decide_equality END TACTIC EXTEND compare -| [ "compare" constr(c1) constr(c2) ] -> [ compare c1 c2 ] +| [ "compare" constr(c1) constr(c2) ] -> [ compare (EConstr.of_constr c1) (EConstr.of_constr c2) ] END diff --git a/proofs/tacmach.ml b/proofs/tacmach.ml index b732e5b9d..aa54e4f9b 100644 --- a/proofs/tacmach.ml +++ b/proofs/tacmach.ml @@ -23,6 +23,8 @@ open Context.Named.Declaration module NamedDecl = Context.Named.Declaration +let compute env sigma c = EConstr.of_constr (compute env sigma c) + let re_sig it gc = { it = it; sigma = gc; } (**************************************************************) diff --git a/proofs/tacmach.mli b/proofs/tacmach.mli index cfbfe12b1..340c7addc 100644 --- a/proofs/tacmach.mli +++ b/proofs/tacmach.mli @@ -69,7 +69,7 @@ val pf_nf : goal sigma -> EConstr.constr -> constr val pf_nf_betaiota : goal sigma -> EConstr.constr -> constr val pf_reduce_to_quantified_ind : goal sigma -> EConstr.types -> pinductive * EConstr.types val pf_reduce_to_atomic_ind : goal sigma -> EConstr.types -> pinductive * EConstr.types -val pf_compute : goal sigma -> EConstr.constr -> constr +val pf_compute : goal sigma -> EConstr.constr -> EConstr.constr val pf_unfoldn : (occurrences * evaluable_global_reference) list -> goal sigma -> EConstr.constr -> constr @@ -127,7 +127,7 @@ module New : sig val pf_hnf_type_of : ('a, 'r) Proofview.Goal.t -> EConstr.constr -> types val pf_whd_all : ('a, 'r) Proofview.Goal.t -> EConstr.t -> constr - val pf_compute : ('a, 'r) Proofview.Goal.t -> EConstr.t -> constr + val pf_compute : ('a, 'r) Proofview.Goal.t -> EConstr.t -> EConstr.constr val pf_matches : ('a, 'r) Proofview.Goal.t -> constr_pattern -> EConstr.constr -> patvar_map diff --git a/tactics/eqdecide.ml b/tactics/eqdecide.ml index be9a34239..a96656d3a 100644 --- a/tactics/eqdecide.ml +++ b/tactics/eqdecide.ml @@ -16,6 +16,7 @@ open Util open Names open Namegen open Term +open EConstr open Declarations open Tactics open Tacticals.New @@ -52,7 +53,6 @@ open Coqlib *) let clear_last = - let open EConstr in Proofview.tclEVARMAP >>= fun sigma -> (onLastHyp (fun c -> (clear [destVar sigma c]))) @@ -70,14 +70,14 @@ let choose_noteq eqonleft = let mkBranches c1 c2 = tclTHENLIST [generalize [c2]; - Simple.elim (EConstr.of_constr c1); + Simple.elim c1; intros; onLastHyp Simple.case; clear_last; intros] let discrHyp id = - let c = { delayed = fun env sigma -> Sigma.here (EConstr.mkVar id, NoBindings) sigma } in + let c = { delayed = fun env sigma -> Sigma.here (mkVar id, NoBindings) sigma } in let tac c = Equality.discr_tac false (Some (None, ElimOnConstr c)) in Tacticals.New.tclDELAYEDWITHHOLES false c tac @@ -89,7 +89,9 @@ let solveNoteqBranch side = (* Constructs the type {c1=c2}+{~c1=c2} *) let make_eq () = -(*FIXME*) Universes.constr_of_global (Coqlib.build_coq_eq ()) +(*FIXME*) EConstr.of_constr (Universes.constr_of_global (Coqlib.build_coq_eq ())) +let build_coq_not () = EConstr.of_constr (build_coq_not ()) +let build_coq_sumbool () = EConstr.of_constr (build_coq_sumbool ()) let mkDecideEqGoal eqonleft op rectype c1 c2 = let equality = mkApp(make_eq(), [|rectype; c1; c2|]) in @@ -116,7 +118,7 @@ let rec rewrite_and_clear hyps = match hyps with | [] -> Proofview.tclUNIT () | id :: hyps -> tclTHENLIST [ - Equality.rewriteLR (EConstr.mkVar id); + Equality.rewriteLR (mkVar id); clear [id]; rewrite_and_clear hyps; ] @@ -125,7 +127,7 @@ let eqCase tac = tclTHEN intro (onLastHypId tac) let injHyp id = - let c = { delayed = fun env sigma -> Sigma.here (EConstr.mkVar id, NoBindings) sigma } in + let c = { delayed = fun env sigma -> Sigma.here (mkVar id, NoBindings) sigma } in let tac c = Equality.injClause None false (Some (None, ElimOnConstr c)) in Tacticals.New.tclDELAYEDWITHHOLES false c tac @@ -137,7 +139,7 @@ let diseqCase hyps eqonleft = (tclTHEN (rewrite_and_clear (List.rev hyps)) (tclTHEN (red_in_concl) (tclTHEN (intro_using absurd) - (tclTHEN (Simple.apply (EConstr.mkVar diseq)) + (tclTHEN (Simple.apply (mkVar diseq)) (tclTHEN (injHyp absurd) (full_trivial [])))))))) @@ -160,9 +162,9 @@ let rec solveArg hyps eqonleft op largs rargs = match largs, rargs with ] | a1 :: largs, a2 :: rargs -> Proofview.Goal.enter { enter = begin fun gl -> - let rectype = pf_unsafe_type_of gl (EConstr.of_constr a1) in + let rectype = pf_unsafe_type_of gl a1 in + let rectype = EConstr.of_constr rectype in let decide = mkDecideEqGoal eqonleft op rectype a1 a2 in - let decide = EConstr.of_constr decide in let tac hyp = solveArg (hyp :: hyps) eqonleft op largs rargs in let subtacs = if eqonleft then [eqCase tac;diseqCase hyps eqonleft;default_auto] @@ -181,7 +183,7 @@ let solveEqBranch rectype = match_eqdec sigma concl >>= fun (eqonleft,op,lhs,rhs,_) -> let (mib,mip) = Global.lookup_inductive rectype in let nparams = mib.mind_nparams in - let getargs l = List.skipn nparams (snd (decompose_app l)) in + let getargs l = List.skipn nparams (snd (decompose_app sigma l)) in let rargs = getargs rhs and largs = getargs lhs in solveArg [] eqonleft op largs rargs @@ -194,7 +196,7 @@ let solveEqBranch rectype = (* The tactic Decide Equality *) -let hd_app c = match kind_of_term c with +let hd_app sigma c = match EConstr.kind sigma c with | App (h,_) -> h | _ -> c @@ -206,13 +208,13 @@ let decideGralEquality = let concl = EConstr.of_constr concl in let sigma = project gl in match_eqdec sigma concl >>= fun (eqonleft,_,c1,c2,typ) -> - let headtyp = hd_app (pf_compute gl (EConstr.of_constr typ)) in - begin match kind_of_term headtyp with + let headtyp = hd_app sigma (pf_compute gl typ) in + begin match EConstr.kind sigma headtyp with | Ind (mi,_) -> Proofview.tclUNIT mi | _ -> tclZEROMSG (Pp.str"This decision procedure only works for inductive objects.") end >>= fun rectype -> (tclTHEN - (mkBranches c1 (EConstr.of_constr c2)) + (mkBranches c1 c2) (tclORELSE (solveNoteqBranch eqonleft) (solveEqBranch rectype))) end } end @@ -227,7 +229,6 @@ let decideEqualityGoal = tclTHEN intros decideGralEquality let decideEquality rectype = Proofview.Goal.enter { enter = begin fun gl -> let decide = mkGenDecideEqGoal rectype gl in - let decide = EConstr.of_constr decide in (tclTHENS (cut decide) [default_auto;decideEqualityGoal]) end } @@ -236,9 +237,9 @@ let decideEquality rectype = let compare c1 c2 = Proofview.Goal.enter { enter = begin fun gl -> - let rectype = pf_unsafe_type_of gl (EConstr.of_constr c1) in + let rectype = pf_unsafe_type_of gl c1 in + let rectype = EConstr.of_constr rectype in let decide = mkDecideEqGoal true (build_coq_sumbool ()) rectype c1 c2 in - let decide = EConstr.of_constr decide in (tclTHENS (cut decide) [(tclTHEN intro (tclTHEN (onLastHyp simplest_case) clear_last)); diff --git a/tactics/eqdecide.mli b/tactics/eqdecide.mli index cb48a5bcc..dca1780b7 100644 --- a/tactics/eqdecide.mli +++ b/tactics/eqdecide.mli @@ -14,4 +14,4 @@ val decideEqualityGoal : unit Proofview.tactic -val compare : Constr.t -> Constr.t -> unit Proofview.tactic +val compare : EConstr.t -> EConstr.t -> unit Proofview.tactic diff --git a/tactics/hipattern.ml b/tactics/hipattern.ml index 36ed589b9..9e78ff323 100644 --- a/tactics/hipattern.ml +++ b/tactics/hipattern.ml @@ -548,7 +548,10 @@ let match_eqdec sigma t = false,op_or,matches sigma (Lazy.force coq_eqdec_rev_pattern) t in match Id.Map.bindings subst with | [(_,typ);(_,c1);(_,c2)] -> - eqonleft, Universes.constr_of_global (Lazy.force op), c1, c2, typ + let typ = EConstr.of_constr typ in + let c1 = EConstr.of_constr c1 in + let c2 = EConstr.of_constr c2 in + eqonleft, EConstr.of_constr (Universes.constr_of_global (Lazy.force op)), c1, c2, typ | _ -> anomaly (Pp.str "Unexpected pattern") (* Patterns "~ ?" and "? -> False" *) diff --git a/tactics/hipattern.mli b/tactics/hipattern.mli index c061c50f0..65ba0aad0 100644 --- a/tactics/hipattern.mli +++ b/tactics/hipattern.mli @@ -143,7 +143,7 @@ val is_matching_sigma : evar_map -> constr -> bool (** Match a decidable equality judgement (e.g [{t=u:>T}+{~t=u}]), returns [t,u,T] and a boolean telling if equality is on the left side *) -val match_eqdec : evar_map -> constr -> bool * Constr.constr * Constr.constr * Constr.constr * Constr.constr +val match_eqdec : evar_map -> constr -> bool * constr * constr * constr * constr (** Match an equality up to conversion; returns [(eq,t1,t2)] in normal form *) val dest_nf_eq : ([ `NF ], 'r) Proofview.Goal.t -> constr -> (constr * constr * constr) -- cgit v1.2.3 From e6a8ab0f428c26fff2bd7e636126974f167101bf Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sun, 20 Nov 2016 01:35:54 +0100 Subject: Tactic_matching API using EConstr. --- dev/top_printers.ml | 2 +- intf/pattern.mli | 4 ++-- ltac/taccoerce.ml | 9 +++++---- ltac/taccoerce.mli | 4 ++-- ltac/tacinterp.ml | 14 +++++++++----- ltac/tactic_matching.ml | 18 ++++++++++-------- ltac/tactic_matching.mli | 8 ++++---- plugins/quote/quote.ml | 2 +- pretyping/constr_matching.ml | 11 +++++------ pretyping/constr_matching.mli | 2 +- pretyping/detyping.ml | 1 + pretyping/patternops.ml | 4 +++- pretyping/pretyping.ml | 2 +- printing/printer.ml | 2 +- tactics/auto.ml | 2 +- tactics/hipattern.ml | 11 +++-------- tactics/tactics.ml | 1 - 17 files changed, 50 insertions(+), 47 deletions(-) diff --git a/dev/top_printers.ml b/dev/top_printers.ml index b7736f498..8290ca945 100644 --- a/dev/top_printers.ml +++ b/dev/top_printers.ml @@ -109,7 +109,7 @@ let ppididmap = ppidmap (fun _ -> pr_id) let prconstrunderbindersidmap = pridmap (fun _ (l,c) -> hov 1 (str"[" ++ prlist_with_sep spc Id.print l ++ str"]") - ++ str "," ++ spc () ++ Termops.print_constr c) + ++ str "," ++ spc () ++ Termops.print_constr (EConstr.Unsafe.to_constr c)) let ppconstrunderbindersidmap l = pp (prconstrunderbindersidmap l) diff --git a/intf/pattern.mli b/intf/pattern.mli index 329ae837e..ac84b91e6 100644 --- a/intf/pattern.mli +++ b/intf/pattern.mli @@ -43,11 +43,11 @@ open Misctypes could be inferred. We also loose the ability of typing ltac variables before calling the right-hand-side of ltac matching clauses. *) -type constr_under_binders = Id.t list * constr +type constr_under_binders = Id.t list * EConstr.constr (** Types of substitutions with or w/o bound variables *) -type patvar_map = constr Id.Map.t +type patvar_map = EConstr.constr Id.Map.t type extended_patvar_map = constr_under_binders Id.Map.t (** {5 Patterns} *) diff --git a/ltac/taccoerce.ml b/ltac/taccoerce.ml index df38a42cb..288e12d0b 100644 --- a/ltac/taccoerce.ml +++ b/ltac/taccoerce.ml @@ -17,7 +17,7 @@ open Geninterp exception CannotCoerceTo of string -let (wit_constr_context : (Empty.t, Empty.t, constr) Genarg.genarg_type) = +let (wit_constr_context : (Empty.t, Empty.t, EConstr.constr) Genarg.genarg_type) = let wit = Genarg.create_arg "constr_context" in let () = register_val0 wit None in wit @@ -64,7 +64,7 @@ let to_constr v = Some c else if has_type v (topwit wit_constr_under_binders) then let vars, c = out_gen (topwit wit_constr_under_binders) v in - match vars with [] -> Some c | _ -> None + match vars with [] -> Some (EConstr.Unsafe.to_constr c) | _ -> None else None let of_uconstr c = in_gen (topwit wit_uconstr) c @@ -96,7 +96,7 @@ let is_variable env id = (* Transforms an id into a constr if possible, or fails with Not_found *) let constr_of_id env id = - Term.mkVar (let _ = Environ.lookup_named id env in id) + EConstr.mkVar (let _ = Environ.lookup_named id env in id) (* Gives the constr corresponding to a Constr_context tactic_arg *) let coerce_to_constr_context v = @@ -212,7 +212,7 @@ let coerce_to_constr env v = | _ -> fail () else if has_type v (topwit wit_constr) then let c = out_gen (topwit wit_constr) v in - ([], c) + ([], EConstr.of_constr c) else if has_type v (topwit wit_constr_under_binders) then out_gen (topwit wit_constr_under_binders) v else if has_type v (topwit wit_var) then @@ -229,6 +229,7 @@ let coerce_to_uconstr env v = let coerce_to_closed_constr env v = let ids,c = coerce_to_constr env v in + let c = EConstr.Unsafe.to_constr c in let () = if not (List.is_empty ids) then raise (CannotCoerceTo "a term") in c diff --git a/ltac/taccoerce.mli b/ltac/taccoerce.mli index 0b67f8726..3049aff7e 100644 --- a/ltac/taccoerce.mli +++ b/ltac/taccoerce.mli @@ -48,7 +48,7 @@ end (** {5 Coercion functions} *) -val coerce_to_constr_context : Value.t -> constr +val coerce_to_constr_context : Value.t -> EConstr.constr val coerce_var_to_ident : bool -> Environ.env -> Value.t -> Id.t @@ -91,6 +91,6 @@ val coerce_to_int_or_var_list : Value.t -> int or_var list (** {5 Missing generic arguments} *) -val wit_constr_context : (Empty.t, Empty.t, constr) genarg_type +val wit_constr_context : (Empty.t, Empty.t, EConstr.constr) genarg_type val wit_constr_under_binders : (Empty.t, Empty.t, constr_under_binders) genarg_type diff --git a/ltac/tacinterp.ml b/ltac/tacinterp.ml index 553565639..d2fb2614e 100644 --- a/ltac/tacinterp.ml +++ b/ltac/tacinterp.ml @@ -239,7 +239,7 @@ let pr_value env v = else if has_type v (topwit wit_constr_context) then let c = out_gen (topwit wit_constr_context) v in match env with - | Some (env,sigma) -> pr_lconstr_env env sigma c + | Some (env,sigma) -> pr_lconstr_env env sigma (EConstr.Unsafe.to_constr c) | _ -> str "a term" else if has_type v (topwit wit_constr) then let c = out_gen (topwit wit_constr) v in @@ -334,7 +334,7 @@ let (+++) lfun1 lfun2 = Id.Map.fold Id.Map.add lfun1 lfun2 let extend_values_with_bindings (ln,lm) lfun = let of_cub c = match c with - | [], c -> Value.of_constr c + | [], c -> Value.of_constr (EConstr.Unsafe.to_constr c) | _ -> in_gen (topwit wit_constr_under_binders) c in (* For compatibility, bound variables are visible only if no other @@ -790,6 +790,7 @@ let interp_may_eval f ist env sigma = function (try let (sigma,ic) = f ist env sigma c in let ctxt = coerce_to_constr_context (Id.Map.find s ist.lfun) in + let ctxt = EConstr.Unsafe.to_constr ctxt in let evdref = ref sigma in let c = subst_meta [Constr_matching.special_meta,ic] ctxt in let c = Typing.e_solve_evars env evdref (EConstr.of_constr c) in @@ -860,6 +861,7 @@ let rec message_of_value v = end } else if has_type v (topwit wit_constr_context) then let c = out_gen (topwit wit_constr_context) v in + let c = EConstr.Unsafe.to_constr c in Ftactic.nf_enter { enter = begin fun gl -> Ftactic.return (pr_constr_env (pf_env gl) (project gl) c) end } else if has_type v (topwit wit_uconstr) then let c = out_gen (topwit wit_uconstr) v in @@ -1464,7 +1466,7 @@ and interp_letin ist llc u = and interp_match_success ist { Tactic_matching.subst ; context ; terms ; lhs } = let (>>=) = Ftactic.bind in let lctxt = Id.Map.map interp_context context in - let hyp_subst = Id.Map.map Value.of_constr terms in + let hyp_subst = Id.Map.map (EConstr.Unsafe.to_constr %> Value.of_constr) terms in let lfun = extend_values_with_bindings subst (lctxt +++ hyp_subst +++ ist.lfun) in let ist = { ist with lfun } in val_interp ist lhs >>= fun v -> @@ -1518,6 +1520,7 @@ and interp_match ist lz constr lmr = Proofview.tclZERO ~info e end end >>= fun constr -> + let constr = EConstr.of_constr constr in Ftactic.enter { enter = begin fun gl -> let sigma = project gl in let env = Proofview.Goal.env gl in @@ -1533,6 +1536,7 @@ and interp_match_goal ist lz lr lmr = let hyps = Proofview.Goal.hyps gl in let hyps = if lr then List.rev hyps else hyps in let concl = Proofview.Goal.concl gl in + let concl = EConstr.of_constr concl in let ilr = read_match_rule (extract_ltac_constr_values ist env) ist env sigma lmr in interp_match_successes lz ist (Tactic_matching.match_goal env sigma hyps concl ilr) end } @@ -1844,7 +1848,7 @@ and interp_atomic ist tac : unit Proofview.tactic = in let c_interp patvars = { Sigma.run = begin fun sigma -> let lfun' = Id.Map.fold (fun id c lfun -> - Id.Map.add id (Value.of_constr c) lfun) + Id.Map.add id (Value.of_constr (EConstr.Unsafe.to_constr c)) lfun) patvars ist.lfun in let sigma = Sigma.to_evar_map sigma in @@ -1871,7 +1875,7 @@ and interp_atomic ist tac : unit Proofview.tactic = let to_catch = function Not_found -> true | e -> CErrors.is_anomaly e in let c_interp patvars = { Sigma.run = begin fun sigma -> let lfun' = Id.Map.fold (fun id c lfun -> - Id.Map.add id (Value.of_constr c) lfun) + Id.Map.add id (Value.of_constr (EConstr.Unsafe.to_constr c)) lfun) patvars ist.lfun in let ist = { ist with lfun = lfun' } in diff --git a/ltac/tactic_matching.ml b/ltac/tactic_matching.ml index 58998c96e..ac64f0bba 100644 --- a/ltac/tactic_matching.ml +++ b/ltac/tactic_matching.ml @@ -23,8 +23,8 @@ module NamedDecl = Context.Named.Declaration substitution mapping corresponding to matched hypotheses. *) type 'a t = { subst : Constr_matching.bound_ident_map * Pattern.extended_patvar_map ; - context : Term.constr Id.Map.t; - terms : Term.constr Id.Map.t; + context : EConstr.constr Id.Map.t; + terms : EConstr.constr Id.Map.t; lhs : 'a; } @@ -84,7 +84,7 @@ let equal_instances env sigma (ctx',c') (ctx,c) = (* How to compare instances? Do we want the terms to be convertible? unifiable? Do we want the universe levels to be relevant? (historically, conv_x is used) *) - CList.equal Id.equal ctx ctx' && Reductionops.is_conv env sigma (EConstr.of_constr c') (EConstr.of_constr c) + CList.equal Id.equal ctx ctx' && Reductionops.is_conv env sigma c' c (** Merges two substitutions. Raises [Not_coherent_metas] when @@ -233,7 +233,7 @@ module PatternMatching (E:StaticEnvironment) = struct | Term p -> begin try - put_subst (Constr_matching.extended_matches E.env E.sigma p (EConstr.of_constr term)) <*> + put_subst (Constr_matching.extended_matches E.env E.sigma p term) <*> return lhs with Constr_matching.PatternMatchingFailure -> fail end @@ -252,7 +252,7 @@ module PatternMatching (E:StaticEnvironment) = struct | Some nctx -> Proofview.tclOR (k lhs nctx) (fun e -> (map s e).stream k ctx) } in - map (Constr_matching.match_subterm_gen E.env E.sigma with_app_context p (EConstr.of_constr term)) imatching_error + map (Constr_matching.match_subterm_gen E.env E.sigma with_app_context p term) imatching_error (** [rule_match_term term rule] matches the term [term] with the @@ -284,8 +284,8 @@ module PatternMatching (E:StaticEnvironment) = struct pick hyps >>= fun decl -> let id = NamedDecl.get_id decl in let refresh = is_local_def decl in - pattern_match_term refresh pat (NamedDecl.get_type decl) () <*> - put_terms (id_map_try_add_name hypname (Term.mkVar id) empty_term_subst) <*> + pattern_match_term refresh pat (EConstr.of_constr (NamedDecl.get_type decl)) () <*> + put_terms (id_map_try_add_name hypname (EConstr.mkVar id) empty_term_subst) <*> return id (** [hyp_match_type hypname bodypat typepat hyps] matches a single @@ -295,9 +295,11 @@ module PatternMatching (E:StaticEnvironment) = struct let hyp_match_body_and_type hypname bodypat typepat hyps = pick hyps >>= function | LocalDef (id,body,hyp) -> + let body = EConstr.of_constr body in + let hyp = EConstr.of_constr hyp in pattern_match_term false bodypat body () <*> pattern_match_term true typepat hyp () <*> - put_terms (id_map_try_add_name hypname (Term.mkVar id) empty_term_subst) <*> + put_terms (id_map_try_add_name hypname (EConstr.mkVar id) empty_term_subst) <*> return id | LocalAssum (id,hyp) -> fail diff --git a/ltac/tactic_matching.mli b/ltac/tactic_matching.mli index 090207bcc..d6f67d6c7 100644 --- a/ltac/tactic_matching.mli +++ b/ltac/tactic_matching.mli @@ -18,8 +18,8 @@ substitution mapping corresponding to matched hypotheses. *) type 'a t = { subst : Constr_matching.bound_ident_map * Pattern.extended_patvar_map ; - context : Term.constr Names.Id.Map.t; - terms : Term.constr Names.Id.Map.t; + context : EConstr.constr Names.Id.Map.t; + terms : EConstr.constr Names.Id.Map.t; lhs : 'a; } @@ -31,7 +31,7 @@ type 'a t = { val match_term : Environ.env -> Evd.evar_map -> - Term.constr -> + EConstr.constr -> (Tacexpr.binding_bound_vars * Pattern.constr_pattern, Tacexpr.glob_tactic_expr) Tacexpr.match_rule list -> Tacexpr.glob_tactic_expr t Proofview.tactic @@ -44,6 +44,6 @@ val match_goal: Environ.env -> Evd.evar_map -> Context.Named.t -> - Term.constr -> + EConstr.constr -> (Tacexpr.binding_bound_vars * Pattern.constr_pattern, Tacexpr.glob_tactic_expr) Tacexpr.match_rule list -> Tacexpr.glob_tactic_expr t Proofview.tactic diff --git a/plugins/quote/quote.ml b/plugins/quote/quote.ml index 2cc402e28..09e77598a 100644 --- a/plugins/quote/quote.ml +++ b/plugins/quote/quote.ml @@ -403,7 +403,7 @@ let quote_terms ivs lc = | (lhs, rhs)::tail -> begin try let s1 = Id.Map.bindings (matches (Global.env ()) Evd.empty rhs (EConstr.of_constr c)) in - let s2 = List.map (fun (i,c_i) -> (coerce_meta_out i,aux c_i)) s1 + let s2 = List.map (fun (i,c_i) -> (coerce_meta_out i,aux (EConstr.Unsafe.to_constr c_i))) s1 in Termops.subst_meta s2 lhs with PatternMatchingFailure -> auxl tail diff --git a/pretyping/constr_matching.ml b/pretyping/constr_matching.ml index 4d2500ccd..06daa5116 100644 --- a/pretyping/constr_matching.ml +++ b/pretyping/constr_matching.ml @@ -16,6 +16,7 @@ open Nameops open Termops open Reductionops open Term +open EConstr open Vars open Pattern open Patternops @@ -61,11 +62,11 @@ let constrain sigma n (ids, m) (names, terms as subst) = let open EConstr in try let (ids', m') = Id.Map.find n terms in - if List.equal Id.equal ids ids' && eq_constr sigma m (EConstr.of_constr m') then subst + if List.equal Id.equal ids ids' && eq_constr sigma m m' then subst else raise PatternMatchingFailure with Not_found -> let () = if Id.Map.mem n names then warn_meta_collision n in - (names, Id.Map.add n (ids, EConstr.Unsafe.to_constr m) terms) + (names, Id.Map.add n (ids, m) terms) let add_binders na1 na2 binding_vars (names, terms as subst) = match na1, na2 with @@ -152,7 +153,6 @@ let merge_binding sigma allow_bound_rels ctx n cT subst = | [] -> (* Optimization *) ([], cT) | _ -> - let open EConstr in let frels = free_rels sigma cT in if allow_bound_rels then let vars = extract_bound_vars frels ctx in @@ -344,7 +344,7 @@ type matching_result = { m_sub : bound_ident_map * patvar_map; m_ctx : constr; } -let mkresult s c n = IStream.Cons ( { m_sub=s; m_ctx=EConstr.Unsafe.to_constr c; } , (IStream.thunk n) ) +let mkresult s c n = IStream.Cons ( { m_sub=s; m_ctx=c; } , (IStream.thunk n) ) let isPMeta = function PMeta _ -> true | _ -> false @@ -362,10 +362,9 @@ let matches_head env sigma pat c = (* Tells if it is an authorized occurrence and if the instance is closed *) let authorized_occ env sigma partial_app closed pat c mk_ctx = - let open EConstr in try let subst = matches_core_closed env sigma false partial_app pat c in - if closed && Id.Map.exists (fun _ c -> not (closed0 c)) (snd subst) + if closed && Id.Map.exists (fun _ c -> not (closed0 sigma c)) (snd subst) then (fun next -> next ()) else (fun next -> mkresult subst (mk_ctx (mkMeta special_meta)) next) with PatternMatchingFailure -> (fun next -> next ()) diff --git a/pretyping/constr_matching.mli b/pretyping/constr_matching.mli index 32bb48c93..4734c90a8 100644 --- a/pretyping/constr_matching.mli +++ b/pretyping/constr_matching.mli @@ -64,7 +64,7 @@ val matches_conv : env -> Evd.evar_map -> constr_pattern -> constr -> patvar_map (whose hole is denoted here with [special_meta]) *) type matching_result = { m_sub : bound_ident_map * patvar_map; - m_ctx : Constr.t } + m_ctx : EConstr.t } (** [match_subterm n pat c] returns the substitution and the context corresponding to each **closed** subterm of [c] matching [pat]. *) diff --git a/pretyping/detyping.ml b/pretyping/detyping.ml index ec8945e85..c0611dcec 100644 --- a/pretyping/detyping.ml +++ b/pretyping/detyping.ml @@ -745,6 +745,7 @@ let detype_closed_glob ?lax isgoal avoid env sigma t = with Not_found -> try (* assumes [detype] does not raise [Not_found] exceptions *) let (b,c) = Id.Map.find id cl.typed in + let c = EConstr.Unsafe.to_constr c in (* spiwack: I'm not sure it is the right thing to do, but I'm computing the detyping environment like [Printer.pr_constr_under_binders_env] does. *) diff --git a/pretyping/patternops.ml b/pretyping/patternops.ml index ffd6e73fa..26e23be23 100644 --- a/pretyping/patternops.ml +++ b/pretyping/patternops.ml @@ -224,6 +224,8 @@ let error_instantiate_pattern id l = ++ strbrk " which " ++ str is ++ strbrk " not bound in the pattern.") let instantiate_pattern env sigma lvar c = + let open EConstr in + let open Vars in let rec aux vars = function | PVar id as x -> (try @@ -235,7 +237,7 @@ let instantiate_pattern env sigma lvar c = ctx in let c = substl inst c in - pattern_of_constr env sigma (EConstr.of_constr c) + pattern_of_constr env sigma c with Not_found (* List.index failed *) -> let vars = List.map_filter (function Name id -> Some id | _ -> None) vars in diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml index 11d50926f..c792bf2ca 100644 --- a/pretyping/pretyping.ml +++ b/pretyping/pretyping.ml @@ -445,7 +445,7 @@ let pretype_id pretype k0 loc env evdref lvar id = try let (ids,c) = Id.Map.find id lvar.ltac_constrs in let subst = List.map (invert_ltac_bound_name lvar env id) ids in - let c = substl subst (EConstr.of_constr c) in + let c = substl subst c in { uj_val = c; uj_type = protected_get_type_of env sigma c } with Not_found -> try let {closure;term} = Id.Map.find id lvar.ltac_uconstrs in diff --git a/printing/printer.ml b/printing/printer.ml index ed6ebdc9b..2a30681bf 100644 --- a/printing/printer.ml +++ b/printing/printer.ml @@ -76,7 +76,7 @@ let pr_constr_under_binders_env_gen pr env sigma (ids,c) = (* we also need to preserve the actual names of the patterns *) (* So what to do? *) let assums = List.map (fun id -> (Name id,(* dummy *) mkProp)) ids in - pr (Termops.push_rels_assum assums env) sigma c + pr (Termops.push_rels_assum assums env) sigma (EConstr.Unsafe.to_constr c) let pr_constr_under_binders_env = pr_constr_under_binders_env_gen pr_constr_env let pr_lconstr_under_binders_env = pr_constr_under_binders_env_gen pr_lconstr_env diff --git a/tactics/auto.ml b/tactics/auto.ml index c34f9dd92..21fe9667b 100644 --- a/tactics/auto.ml +++ b/tactics/auto.ml @@ -152,7 +152,7 @@ let conclPattern concl pat tac = let open Genarg in let open Geninterp in let inj c = match val_tag (topwit Stdarg.wit_constr) with - | Val.Base tag -> Val.Dyn (tag, c) + | Val.Base tag -> Val.Dyn (tag, EConstr.Unsafe.to_constr c) | _ -> assert false in let fold id c accu = Id.Map.add id (inj c) accu in diff --git a/tactics/hipattern.ml b/tactics/hipattern.ml index 9e78ff323..b92d65908 100644 --- a/tactics/hipattern.ml +++ b/tactics/hipattern.ml @@ -353,8 +353,6 @@ let is_imp_term sigma c = op2bool (match_with_imp_term sigma c) let match_with_nottype sigma t = try let (arg,mind) = match_arrow_pattern sigma t in - let arg = EConstr.of_constr arg in - let mind = EConstr.of_constr mind in if is_empty_type sigma mind then Some (mind,arg) else None with PatternMatchingFailure -> None @@ -470,11 +468,11 @@ let match_eq_nf gls eqn (ref, hetero) = let n = if hetero then 4 else 3 in let args = List.init n (fun i -> mkGPatVar ("X" ^ string_of_int (i + 1))) in let pat = mkPattern (mkGAppRef ref args) in - let pf_whd_all gls c = EConstr.of_constr (pf_whd_all gls (EConstr.of_constr c)) in + let pf_whd_all gls c = EConstr.of_constr (pf_whd_all gls c) in match Id.Map.bindings (pf_matches gls pat eqn) with | [(m1,t);(m2,x);(m3,y)] -> assert (Id.equal m1 meta1 && Id.equal m2 meta2 && Id.equal m3 meta3); - (EConstr.of_constr t,pf_whd_all gls x,pf_whd_all gls y) + (t,pf_whd_all gls x,pf_whd_all gls y) | _ -> anomaly ~label:"match_eq" (Pp.str "an eq pattern should match 3 terms") let dest_nf_eq gls eqn = @@ -502,7 +500,7 @@ let coq_sig_pattern = let match_sigma sigma t = match Id.Map.bindings (matches sigma (Lazy.force coq_sig_pattern) t) with - | [(_,a); (_,p)] -> (EConstr.of_constr a,EConstr.of_constr p) + | [(_,a); (_,p)] -> (a,p) | _ -> anomaly (Pp.str "Unexpected pattern") let is_matching_sigma sigma t = is_matching sigma (Lazy.force coq_sig_pattern) t @@ -548,9 +546,6 @@ let match_eqdec sigma t = false,op_or,matches sigma (Lazy.force coq_eqdec_rev_pattern) t in match Id.Map.bindings subst with | [(_,typ);(_,c1);(_,c2)] -> - let typ = EConstr.of_constr typ in - let c1 = EConstr.of_constr c1 in - let c2 = EConstr.of_constr c2 in eqonleft, EConstr.of_constr (Universes.constr_of_global (Lazy.force op)), c1, c2, typ | _ -> anomaly (Pp.str "Unexpected pattern") diff --git a/tactics/tactics.ml b/tactics/tactics.ml index c025ca9b5..606eb27b9 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -816,7 +816,6 @@ let e_change_in_hyp redfun (id,where) = type change_arg = Pattern.patvar_map -> EConstr.constr Sigma.run let make_change_arg c pats = - let pats = Id.Map.map EConstr.of_constr pats in { run = fun sigma -> Sigma.here (replace_vars (Id.Map.bindings pats) c) sigma } let check_types env sigma mayneedglobalcheck deep newc origc = -- cgit v1.2.3 From 53978e8dd1cc1f1bb7581ea33cee7b9d870ed418 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sun, 20 Nov 2016 02:22:05 +0100 Subject: Extratactics API using EConstr. --- ltac/extratactics.ml4 | 115 ++++++++++++++++++++++++++++++-------------------- 1 file changed, 70 insertions(+), 45 deletions(-) diff --git a/ltac/extratactics.ml4 b/ltac/extratactics.ml4 index f87aa407d..d53248a04 100644 --- a/ltac/extratactics.ml4 +++ b/ltac/extratactics.ml4 @@ -340,6 +340,9 @@ END (**********************************************************************) (* Refine *) +open EConstr +open Vars + let constr_flags = { Pretyping.use_typeclasses = true; Pretyping.use_unif_heuristics = true; @@ -396,8 +399,6 @@ VERNAC COMMAND EXTEND DeriveInversionClear -> [ add_inversion_lemma_exn na c GProp false inv_clear_tac ] END -open Term - VERNAC COMMAND EXTEND DeriveInversion | [ "Derive" "Inversion" ident(na) "with" constr(c) "Sort" sort(s) ] => [ seff na ] @@ -494,7 +495,7 @@ let cache_transitivity_lemma (_,(left,lem)) = let subst_transitivity_lemma (subst,(b,ref)) = (b,subst_mps subst ref) -let inTransitivity : bool * constr -> obj = +let inTransitivity : bool * Constr.constr -> obj = declare_object {(default_object "TRANSITIVITY-STEPS") with cache_function = cache_transitivity_lemma; open_function = (fun i o -> if Int.equal i 1 then cache_transitivity_lemma o); @@ -646,6 +647,7 @@ let hResolve id c occ t = let sigma = Sigma.to_evar_map sigma in let env = Termops.clear_named_body id (Proofview.Goal.env gl) in let concl = Proofview.Goal.concl gl in + let concl = EConstr.of_constr concl in let env_ids = Termops.ids_of_context env in let c_raw = Detyping.detype true env_ids env sigma c in let t_raw = Detyping.detype true env_ids env sigma t in @@ -659,10 +661,12 @@ let hResolve id c occ t = resolve_hole (subst_hole_with_term (fst (Loc.unloc loc)) c_raw t_hole) in let t_constr,ctx = resolve_hole (subst_var_with_hole occ id t_raw) in + let t_constr = EConstr.of_constr t_constr in let sigma = Evd.merge_universe_context sigma ctx in - let t_constr_type = Retyping.get_type_of env sigma (EConstr.of_constr t_constr) in + let t_constr_type = Retyping.get_type_of env sigma t_constr in + let t_constr_type = EConstr.of_constr t_constr_type in let tac = - (change_concl (EConstr.of_constr (mkLetIn (Anonymous,t_constr,t_constr_type,concl)))) + (change_concl (mkLetIn (Anonymous,t_constr,t_constr_type,concl))) in Sigma.Unsafe.of_pair (tac, sigma) end } @@ -691,12 +695,14 @@ let hget_evar n = let sigma = Tacmach.New.project gl in let concl = Proofview.Goal.concl gl in let evl = evar_list concl in + let concl = EConstr.of_constr concl in if List.length evl < n then error "Not enough uninstantiated existential variables."; if n <= 0 then error "Incorrect existential variable index."; let ev = List.nth evl (n-1) in + let ev = (fst ev, Array.map EConstr.of_constr (snd ev)) in let ev_type = existential_type sigma ev in - change_concl (EConstr.of_constr (mkLetIn (Anonymous,mkEvar ev,ev_type,concl))) + change_concl (mkLetIn (Anonymous,mkEvar ev,ev_type,concl)) end } TACTIC EXTEND hget_evar @@ -719,7 +725,7 @@ let rewrite_except h = Proofview.Goal.nf_enter { enter = begin fun gl -> let hyps = Tacmach.New.pf_ids_of_hyps gl in Tacticals.New.tclMAP (fun id -> if Id.equal id h then Proofview.tclUNIT () else - Tacticals.New.tclTRY (Equality.general_rewrite_in true Locus.AllOccurrences true true id (EConstr.mkVar h) false)) + Tacticals.New.tclTRY (Equality.general_rewrite_in true Locus.AllOccurrences true true id (mkVar h) false)) hyps end } @@ -736,18 +742,20 @@ let refl_equal = call it before it is defined. *) let mkCaseEq a : unit Proofview.tactic = Proofview.Goal.nf_enter { enter = begin fun gl -> - let type_of_a = Tacmach.New.of_old (fun g -> Tacmach.pf_unsafe_type_of g (EConstr.of_constr a)) gl in + let type_of_a = Tacmach.New.of_old (fun g -> Tacmach.pf_unsafe_type_of g a) gl in + let type_of_a = EConstr.of_constr type_of_a in Tacticals.New.tclTHENLIST - [Tactics.generalize [EConstr.of_constr (mkApp(delayed_force refl_equal, [| type_of_a; a|]))]; + [Tactics.generalize [(mkApp(EConstr.of_constr (delayed_force refl_equal), [| type_of_a; a|]))]; Proofview.Goal.nf_enter { enter = begin fun gl -> let concl = Proofview.Goal.concl gl in + let concl = EConstr.of_constr concl in let env = Proofview.Goal.env gl in (** FIXME: this looks really wrong. Does anybody really use this tactic? *) - let Sigma (c, _, _) = (Tacred.pattern_occs [Locus.OnlyOccurrences [1], EConstr.of_constr a]).Reductionops.e_redfun env (Sigma.Unsafe.of_evar_map Evd.empty) (EConstr.of_constr concl) in + let Sigma (c, _, _) = (Tacred.pattern_occs [Locus.OnlyOccurrences [1], a]).Reductionops.e_redfun env (Sigma.Unsafe.of_evar_map Evd.empty) concl in let c = EConstr.of_constr c in change_concl c end }; - simplest_case (EConstr.of_constr a)] + simplest_case a] end } @@ -770,36 +778,38 @@ let case_eq_intros_rewrite x = ] end } -let rec find_a_destructable_match t = +let rec find_a_destructable_match sigma t = let cl = induction_arg_of_quantified_hyp (NamedHyp (Id.of_string "x")) in let cl = [cl, (None, None), None], None in let dest = TacAtom (Loc.ghost, TacInductionDestruct(false, false, cl)) in - match kind_of_term t with - | Case (_,_,x,_) when closed0 x -> - if isVar x then + match EConstr.kind sigma t with + | Case (_,_,x,_) when closed0 sigma x -> + if isVar sigma x then (* TODO check there is no rel n. *) raise (Found (Tacinterp.eval_tactic dest)) else (* let _ = Pp.msgnl (Printer.pr_lconstr x) in *) raise (Found (case_eq_intros_rewrite x)) - | _ -> iter_constr find_a_destructable_match t + | _ -> EConstr.iter sigma (fun c -> find_a_destructable_match sigma c) t let destauto t = - try find_a_destructable_match t; + Proofview.tclEVARMAP >>= fun sigma -> + try find_a_destructable_match sigma t; Tacticals.New.tclZEROMSG (str "No destructable match found") with Found tac -> tac let destauto_in id = Proofview.Goal.nf_enter { enter = begin fun gl -> - let ctype = Tacmach.New.of_old (fun g -> Tacmach.pf_unsafe_type_of g (EConstr.mkVar id)) gl in + let ctype = Tacmach.New.of_old (fun g -> Tacmach.pf_unsafe_type_of g (mkVar id)) gl in + let ctype = EConstr.of_constr ctype in (* Pp.msgnl (Printer.pr_lconstr (mkVar id)); *) (* Pp.msgnl (Printer.pr_lconstr (ctype)); *) destauto ctype end } TACTIC EXTEND destauto -| [ "destauto" ] -> [ Proofview.Goal.nf_enter { enter = begin fun gl -> destauto (Proofview.Goal.concl gl) end } ] +| [ "destauto" ] -> [ Proofview.Goal.nf_enter { enter = begin fun gl -> destauto (EConstr.of_constr (Proofview.Goal.concl gl)) end } ] | [ "destauto" "in" hyp(id) ] -> [ destauto_in id ] END @@ -819,21 +829,24 @@ END TACTIC EXTEND constr_eq_nounivs | [ "constr_eq_nounivs" constr(x) constr(y) ] -> [ - if eq_constr_nounivs x y then Proofview.tclUNIT () else Tacticals.New.tclFAIL 0 (str "Not equal") ] + Proofview.tclEVARMAP >>= fun sigma -> + let x = EConstr.of_constr x in + let y = EConstr.of_constr y in + if eq_constr_nounivs sigma x y then Proofview.tclUNIT () else Tacticals.New.tclFAIL 0 (str "Not equal") ] END TACTIC EXTEND is_evar -| [ "is_evar" constr(x) ] -> - [ Proofview.tclBIND Proofview.tclEVARMAP begin fun sigma -> - match Evarutil.kind_of_term_upto sigma x with - | Evar _ -> Proofview.tclUNIT () - | _ -> Tacticals.New.tclFAIL 0 (str "Not an evar") - end +| [ "is_evar" constr(x) ] -> [ + Proofview.tclEVARMAP >>= fun sigma -> + match EConstr.kind sigma (EConstr.of_constr x) with + | Evar _ -> Proofview.tclUNIT () + | _ -> Tacticals.New.tclFAIL 0 (str "Not an evar") ] END +let has_evar sigma c = let rec has_evar x = - match kind_of_term x with + match EConstr.kind sigma x with | Evar _ -> true | Rel _ | Var _ | Meta _ | Sort _ | Const _ | Ind _ | Construct _ -> false @@ -852,57 +865,68 @@ and has_evar_array x = Array.exists has_evar x and has_evar_prec (_, ts1, ts2) = Array.exists has_evar ts1 || Array.exists has_evar ts2 +in +has_evar c TACTIC EXTEND has_evar -| [ "has_evar" constr(x) ] -> - [ if has_evar x then Proofview.tclUNIT () else Tacticals.New.tclFAIL 0 (str "No evars") ] +| [ "has_evar" constr(x) ] -> [ + Proofview.tclEVARMAP >>= fun sigma -> + if has_evar sigma (EConstr.of_constr x) then Proofview.tclUNIT () else Tacticals.New.tclFAIL 0 (str "No evars") +] END TACTIC EXTEND is_hyp -| [ "is_var" constr(x) ] -> - [ match kind_of_term x with +| [ "is_var" constr(x) ] -> [ + Proofview.tclEVARMAP >>= fun sigma -> + match EConstr.kind sigma (EConstr.of_constr x) with | Var _ -> Proofview.tclUNIT () | _ -> Tacticals.New.tclFAIL 0 (str "Not a variable or hypothesis") ] END TACTIC EXTEND is_fix -| [ "is_fix" constr(x) ] -> - [ match kind_of_term x with +| [ "is_fix" constr(x) ] -> [ + Proofview.tclEVARMAP >>= fun sigma -> + match EConstr.kind sigma (EConstr.of_constr x) with | Fix _ -> Proofview.tclUNIT () | _ -> Tacticals.New.tclFAIL 0 (Pp.str "not a fix definition") ] END;; TACTIC EXTEND is_cofix -| [ "is_cofix" constr(x) ] -> - [ match kind_of_term x with +| [ "is_cofix" constr(x) ] -> [ + Proofview.tclEVARMAP >>= fun sigma -> + match EConstr.kind sigma (EConstr.of_constr x) with | CoFix _ -> Proofview.tclUNIT () | _ -> Tacticals.New.tclFAIL 0 (Pp.str "not a cofix definition") ] END;; TACTIC EXTEND is_ind -| [ "is_ind" constr(x) ] -> - [ match kind_of_term x with +| [ "is_ind" constr(x) ] -> [ + Proofview.tclEVARMAP >>= fun sigma -> + match EConstr.kind sigma (EConstr.of_constr x) with | Ind _ -> Proofview.tclUNIT () | _ -> Tacticals.New.tclFAIL 0 (Pp.str "not an (co)inductive datatype") ] END;; TACTIC EXTEND is_constructor -| [ "is_constructor" constr(x) ] -> - [ match kind_of_term x with +| [ "is_constructor" constr(x) ] -> [ + Proofview.tclEVARMAP >>= fun sigma -> + match EConstr.kind sigma (EConstr.of_constr x) with | Construct _ -> Proofview.tclUNIT () | _ -> Tacticals.New.tclFAIL 0 (Pp.str "not a constructor") ] END;; TACTIC EXTEND is_proj -| [ "is_proj" constr(x) ] -> - [ match kind_of_term x with +| [ "is_proj" constr(x) ] -> [ + Proofview.tclEVARMAP >>= fun sigma -> + match EConstr.kind sigma (EConstr.of_constr x) with | Proj _ -> Proofview.tclUNIT () | _ -> Tacticals.New.tclFAIL 0 (Pp.str "not a primitive projection") ] END;; TACTIC EXTEND is_const -| [ "is_const" constr(x) ] -> - [ match kind_of_term x with +| [ "is_const" constr(x) ] -> [ + Proofview.tclEVARMAP >>= fun sigma -> + match EConstr.kind sigma (EConstr.of_constr x) with | Const _ -> Proofview.tclUNIT () | _ -> Tacticals.New.tclFAIL 0 (Pp.str "not a constant") ] END;; @@ -1046,8 +1070,9 @@ END let decompose l c = Proofview.Goal.enter { enter = begin fun gl -> + let sigma = Tacmach.New.project gl in let to_ind c = - if isInd c then Univ.out_punivs (destInd c) + if isInd sigma c then Univ.out_punivs (destInd sigma c) else error "not an inductive type" in let l = List.map to_ind l in @@ -1055,7 +1080,7 @@ let decompose l c = end } TACTIC EXTEND decompose -| [ "decompose" "[" ne_constr_list(l) "]" constr(c) ] -> [ decompose l (EConstr.of_constr c) ] +| [ "decompose" "[" ne_constr_list(l) "]" constr(c) ] -> [ decompose (List.map EConstr.of_constr l) (EConstr.of_constr c) ] END (** library/keys *) -- cgit v1.2.3 From f6f271d3179e6d8acd90cbd07c8e60e5a68efb17 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sun, 20 Nov 2016 02:46:50 +0100 Subject: G_auto API using EConstr. --- ltac/g_auto.ml4 | 4 +++- 1 file changed, 3 insertions(+), 1 deletion(-) diff --git a/ltac/g_auto.ml4 b/ltac/g_auto.ml4 index f540beb56..b5e56d93b 100644 --- a/ltac/g_auto.ml4 +++ b/ltac/g_auto.ml4 @@ -16,6 +16,7 @@ open Pcoq.Constr open Pltac open Hints open Tacexpr +open Proofview.Notations DECLARE PLUGIN "g_auto" @@ -149,7 +150,8 @@ TACTIC EXTEND autounfold_one TACTIC EXTEND autounfoldify | [ "autounfoldify" constr(x) ] -> [ - let db = match Term.kind_of_term x with + Proofview.tclEVARMAP >>= fun sigma -> + let db = match EConstr.kind sigma (EConstr.of_constr x) with | Term.Const (c,_) -> Names.Label.to_string (Names.con_label c) | _ -> assert false in Eauto.autounfold ["core";db] Locusops.onConcl -- cgit v1.2.3 From c0d38ae52ac410811e7df54c52e8d3a18bb11bcb Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sun, 20 Nov 2016 02:50:07 +0100 Subject: G_class API using Econstr. --- ltac/g_class.ml4 | 12 ++++++++---- 1 file changed, 8 insertions(+), 4 deletions(-) diff --git a/ltac/g_class.ml4 b/ltac/g_class.ml4 index 45ee86c4e..d4b84284f 100644 --- a/ltac/g_class.ml4 +++ b/ltac/g_class.ml4 @@ -83,18 +83,22 @@ open Term open Proofview.Goal open Proofview.Notations -let rec eq_constr_mod_evars x y = - match kind_of_term x, kind_of_term y with +let rec eq_constr_mod_evars sigma x y = + let open EConstr in + match EConstr.kind sigma x, EConstr.kind sigma y with | Evar (e1, l1), Evar (e2, l2) when not (Evar.equal e1 e2) -> true - | _, _ -> compare_constr eq_constr_mod_evars x y + | _, _ -> compare_constr sigma (fun x y -> eq_constr_mod_evars sigma x y) x y let progress_evars t = Proofview.Goal.nf_enter { enter = begin fun gl -> let concl = Proofview.Goal.concl gl in + let concl = EConstr.of_constr concl in let check = Proofview.Goal.nf_enter { enter = begin fun gl' -> + let sigma = Tacmach.New.project gl' in let newconcl = Proofview.Goal.concl gl' in - if eq_constr_mod_evars concl newconcl + let newconcl = EConstr.of_constr newconcl in + if eq_constr_mod_evars sigma concl newconcl then Tacticals.New.tclFAIL 0 (Pp.str"No progress made (modulo evars)") else Proofview.tclUNIT () end } -- cgit v1.2.3 From d833b81b49366e95cf20a1d00f9c63883adb8942 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sun, 20 Nov 2016 03:04:13 +0100 Subject: Rewrite API using EConstr. --- engine/evarutil.ml | 3 +- engine/evarutil.mli | 4 +- engine/termops.ml | 5 + engine/termops.mli | 2 + ltac/rewrite.ml | 378 ++++++++++++++++++++-------------- ltac/rewrite.mli | 7 +- ltac/tacinterp.ml | 3 +- ltac/tacinterp.mli | 2 +- plugins/decl_mode/decl_proof_instr.ml | 7 +- plugins/setoid_ring/newring.ml | 29 ++- pretyping/reductionops.ml | 3 +- pretyping/reductionops.mli | 2 +- pretyping/unification.ml | 10 +- proofs/clenv.ml | 2 +- 14 files changed, 276 insertions(+), 181 deletions(-) diff --git a/engine/evarutil.ml b/engine/evarutil.ml index 8b75d8242..204997445 100644 --- a/engine/evarutil.ml +++ b/engine/evarutil.ml @@ -463,12 +463,13 @@ let e_new_type_evar env evdref ?src ?filter ?naming ?principal rigid = c let new_Type ?(rigid=Evd.univ_flexible) env evd = + let open EConstr in let Sigma (s, sigma, p) = Sigma.new_sort_variable rigid evd in Sigma (mkSort s, sigma, p) 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 = diff --git a/engine/evarutil.mli b/engine/evarutil.mli index 1b592b790..cf36f5953 100644 --- a/engine/evarutil.mli +++ b/engine/evarutil.mli @@ -52,8 +52,8 @@ val e_new_type_evar : env -> evar_map ref -> ?src:Loc.t * Evar_kinds.t -> ?filter:Filter.t -> ?naming:Misctypes.intro_pattern_naming_expr -> ?principal:bool -> rigid -> EConstr.constr * sorts -val new_Type : ?rigid:rigid -> env -> 'r Sigma.t -> (constr, 'r) Sigma.sigma -val e_new_Type : ?rigid:rigid -> env -> evar_map ref -> constr +val new_Type : ?rigid:rigid -> env -> 'r Sigma.t -> (EConstr.constr, 'r) Sigma.sigma +val e_new_Type : ?rigid:rigid -> env -> evar_map ref -> EConstr.constr val restrict_evar : 'r Sigma.t -> existential_key -> Filter.t -> constr list option -> (existential_key, 'r) Sigma.sigma diff --git a/engine/termops.ml b/engine/termops.ml index 7c89f190f..ecc6ab68e 100644 --- a/engine/termops.ml +++ b/engine/termops.ml @@ -858,6 +858,11 @@ 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 (Prop Null) -> true + | Cast (c,_,_) -> is_Prop sigma c + | _ -> false + (* eq_constr extended with universe erasure *) let compare_constr_univ sigma f cv_pb t1 t2 = match EConstr.kind sigma t1, EConstr.kind sigma t2 with diff --git a/engine/termops.mli b/engine/termops.mli index a865c80fb..05604beda 100644 --- a/engine/termops.mli +++ b/engine/termops.mli @@ -275,6 +275,8 @@ val isGlobalRef : Evd.evar_map -> EConstr.t -> bool val is_template_polymorphic : env -> Evd.evar_map -> EConstr.t -> bool +val is_Prop : Evd.evar_map -> EConstr.t -> bool + (** Combinators on judgments *) val on_judgment : ('a -> 'b) -> ('a, 'a) punsafe_judgment -> ('b, 'b) punsafe_judgment diff --git a/ltac/rewrite.ml b/ltac/rewrite.ml index a2b6cb97c..ecb653587 100644 --- a/ltac/rewrite.ml +++ b/ltac/rewrite.ml @@ -13,6 +13,7 @@ open Util open Nameops open Namegen open Term +open EConstr open Vars open Reduction open Tacticals.New @@ -31,6 +32,7 @@ open Decl_kinds open Elimschemes open Environ open Termops +open EConstr open Libnames open Sigma.Notations open Proofview.Notations @@ -43,6 +45,21 @@ let local_assum (na, t) = let inj = EConstr.Unsafe.to_constr in RelDecl.LocalAssum (na, inj t) +let local_def (na, b, t) = + let inj = EConstr.Unsafe.to_constr in + RelDecl.LocalDef (na, inj b, inj t) + +let nlocal_assum (na, t) = + let inj = EConstr.Unsafe.to_constr in + NamedDecl.LocalAssum (na, inj t) + +let nlocal_def (na, b, t) = + let inj = EConstr.Unsafe.to_constr in + NamedDecl.LocalDef (na, inj b, inj t) + +let nf_evar sigma c = + EConstr.of_constr (Evarutil.nf_evar sigma (EConstr.Unsafe.to_constr c)) + (** Typeclass-based generalized rewriting. *) (** Constants used by the tactic. *) @@ -77,6 +94,7 @@ let find_global dir s = fun (evd,cstrs) -> let sigma = Sigma.Unsafe.of_evar_map evd in let Sigma (c, sigma, _) = Evarutil.new_global sigma (Lazy.force gr) in + let c = EConstr.of_constr c in let evd = Sigma.to_evar_map sigma in (evd, cstrs), c @@ -99,10 +117,10 @@ let cstrevars evars = snd evars let new_cstr_evar (evd,cstrs) env t = let s = Typeclasses.set_resolvable Evd.Store.empty false in let evd = Sigma.Unsafe.of_evar_map evd in - let Sigma (t, evd', _) = Evarutil.new_evar ~store:s env evd (EConstr.of_constr t) in + let Sigma (t, evd', _) = Evarutil.new_evar ~store:s env evd t in let evd' = Sigma.to_evar_map evd' in - let ev, _ = EConstr.destEvar evd' t in - (evd', Evar.Set.add ev cstrs), EConstr.Unsafe.to_constr t + let ev, _ = destEvar evd' t in + (evd', Evar.Set.add ev cstrs), t (** Building or looking up instances. *) let e_new_cstr_evar env evars t = @@ -117,7 +135,8 @@ let extends_undefined evars evars' = let app_poly_check env evars f args = let (evars, cstrs), fc = f evars in let evdref = ref evars in - let t = Typing.e_solve_evars env evdref (EConstr.of_constr (mkApp (fc, args))) in + let t = Typing.e_solve_evars env evdref (mkApp (fc, args)) in + let t = EConstr.of_constr t in (!evdref, cstrs), t let app_poly_nocheck env evars f args = @@ -131,8 +150,7 @@ let app_poly_sort b = let find_class_proof proof_type proof_method env evars carrier relation = try let evars, goal = app_poly_check env evars proof_type [| carrier ; relation |] in - let evars', c = Typeclasses.resolve_one_typeclass env (goalevars evars) (EConstr.of_constr goal) in - let c = EConstr.Unsafe.to_constr c in + let evars', c = Typeclasses.resolve_one_typeclass env (goalevars evars) goal in if extends_undefined (goalevars evars) evars' then raise Not_found else app_poly_check env (evars',cstrevars evars) proof_method [| carrier; relation; c |] with e when Logic.catchable_exception e -> raise Not_found @@ -188,6 +206,7 @@ end) = struct fun (evd,cstrs) -> let sigma = Sigma.Unsafe.of_evar_map evd in let Sigma (c, sigma, _) = Evarutil.new_global sigma (Lazy.force l) in + let c = EConstr.of_constr c in let evd = Sigma.to_evar_map sigma in (evd, cstrs), c @@ -196,6 +215,7 @@ end) = struct fun (evd,cstrs) -> let sigma = Sigma.Unsafe.of_evar_map evd in let Sigma (c, sigma, _) = Evarutil.new_global sigma (Lazy.force l) in + let c = EConstr.of_constr c in let evd = Sigma.to_evar_map sigma in (evd, cstrs), c @@ -219,28 +239,32 @@ end) = struct match obj with | None | Some (_, None) -> let evars, relty = mk_relation env evars ty in - if closed0 ty then + if closed0 (goalevars evars) ty then let env' = Environ.reset_with_named_context (Environ.named_context_val env) env in new_cstr_evar evars env' relty else new_cstr_evar evars newenv relty | Some (x, Some rel) -> evars, rel in let rec aux env evars ty l = - let t = Reductionops.whd_all env (goalevars evars) (EConstr.of_constr ty) in - match kind_of_term t, l with + let t = Reductionops.whd_all env (goalevars evars) ty in + let t = EConstr.of_constr t in + match EConstr.kind (goalevars evars) t, l with | Prod (na, ty, b), obj :: cstrs -> - let b = Reductionops.nf_betaiota (goalevars evars) (EConstr.of_constr b) in - if noccurn 1 b (* non-dependent product *) then - let ty = Reductionops.nf_betaiota (goalevars evars) (EConstr.of_constr ty) in + let b = Reductionops.nf_betaiota (goalevars evars) b in + let b = EConstr.of_constr b in + if noccurn (goalevars evars) 1 b (* non-dependent product *) then + let ty = Reductionops.nf_betaiota (goalevars evars) ty in + let ty = EConstr.of_constr ty in let (evars, b', arg, cstrs) = aux env evars (subst1 mkProp b) cstrs in let evars, relty = mk_relty evars env ty obj in let evars, newarg = app_poly env evars respectful [| ty ; b' ; relty ; arg |] in evars, mkProd(na, ty, b), newarg, (ty, Some relty) :: cstrs else let (evars, b, arg, cstrs) = - aux (Environ.push_rel (LocalAssum (na, ty)) env) evars b cstrs + aux (Environ.push_rel (local_assum (na, ty)) env) evars b cstrs in - let ty = Reductionops.nf_betaiota (goalevars evars) (EConstr.of_constr ty) in + let ty = Reductionops.nf_betaiota (goalevars evars) ty in + let ty = EConstr.of_constr ty in let pred = mkLambda (na, ty, b) in let liftarg = mkLambda (na, ty, arg) in let evars, arg' = app_poly env evars forall_relation [| ty ; pred ; liftarg |] in @@ -250,7 +274,8 @@ end) = struct | _, [] -> (match finalcstr with | None | Some (_, None) -> - let t = Reductionops.nf_betaiota (fst evars) (EConstr.of_constr ty) in + let t = Reductionops.nf_betaiota (fst evars) ty in + let t = EConstr.of_constr t in let evars, rel = mk_relty evars env t None in evars, t, rel, [t, Some rel] | Some (t, Some rel) -> evars, t, rel, [t, Some rel]) @@ -258,30 +283,30 @@ end) = struct (** Folding/unfolding of the tactic constants. *) - let unfold_impl t = - match kind_of_term t with + let unfold_impl sigma t = + match EConstr.kind sigma t with | App (arrow, [| a; b |])(* when eq_constr arrow (Lazy.force impl) *) -> mkProd (Anonymous, a, lift 1 b) | _ -> assert false - let unfold_all t = - match kind_of_term t with + let unfold_all sigma t = + match EConstr.kind sigma t with | App (id, [| a; b |]) (* when eq_constr id (Lazy.force coq_all) *) -> - (match kind_of_term b with + (match EConstr.kind sigma b with | Lambda (n, ty, b) -> mkProd (n, ty, b) | _ -> assert false) | _ -> assert false - let unfold_forall t = - match kind_of_term t with + let unfold_forall sigma t = + match EConstr.kind sigma t with | App (id, [| a; b |]) (* when eq_constr id (Lazy.force coq_all) *) -> - (match kind_of_term b with + (match EConstr.kind sigma b with | Lambda (n, ty, b) -> mkProd (n, ty, b) | _ -> assert false) | _ -> assert false let arrow_morphism env evd ta tb a b = - let ap = is_Prop ta and bp = is_Prop tb in + let ap = is_Prop (goalevars evd) ta and bp = is_Prop (goalevars evd) tb in if ap && bp then app_poly env evd impl [| a; b |], unfold_impl else if ap then (* Domain in Prop, CoDomain in Type *) (app_poly env evd arrow [| a; b |]), unfold_impl @@ -294,26 +319,25 @@ end) = struct let rec decomp_pointwise sigma n c = if Int.equal n 0 then c else - let open EConstr in - match kind_of_term c with - | App (f, [| a; b; relb |]) when Globnames.is_global (pointwise_relation_ref ()) f -> + match EConstr.kind sigma c with + | App (f, [| a; b; relb |]) when Termops.is_global sigma (pointwise_relation_ref ()) f -> decomp_pointwise sigma (pred n) relb - | App (f, [| a; b; arelb |]) when Globnames.is_global (forall_relation_ref ()) f -> - decomp_pointwise sigma (pred n) (Reductionops.beta_applist sigma (EConstr.of_constr arelb, [mkRel 1])) + | App (f, [| a; b; arelb |]) when Termops.is_global sigma (forall_relation_ref ()) f -> + decomp_pointwise sigma (pred n) (EConstr.of_constr (Reductionops.beta_applist sigma (arelb, [mkRel 1]))) | _ -> invalid_arg "decomp_pointwise" let rec apply_pointwise sigma rel = function | arg :: args -> - (match kind_of_term rel with - | App (f, [| a; b; relb |]) when Globnames.is_global (pointwise_relation_ref ()) f -> + (match EConstr.kind sigma rel with + | App (f, [| a; b; relb |]) when Termops.is_global sigma (pointwise_relation_ref ()) f -> apply_pointwise sigma relb args - | App (f, [| a; b; arelb |]) when Globnames.is_global (forall_relation_ref ()) f -> - apply_pointwise sigma (Reductionops.beta_applist sigma (EConstr.of_constr arelb, [EConstr.of_constr arg])) args + | App (f, [| a; b; arelb |]) when Termops.is_global sigma (forall_relation_ref ()) f -> + apply_pointwise sigma (EConstr.of_constr (Reductionops.beta_applist sigma (arelb, [arg]))) args | _ -> invalid_arg "apply_pointwise") | [] -> rel let pointwise_or_dep_relation env evd n t car rel = - if noccurn 1 car && noccurn 1 rel then + if noccurn (goalevars evd) 1 car && noccurn (goalevars evd) 1 rel then app_poly env evd pointwise_relation [| t; lift (-1) car; lift (-1) rel |] else app_poly env evd forall_relation @@ -330,14 +354,15 @@ end) = struct let rec aux evars env prod n = if Int.equal n 0 then start evars env prod else - match kind_of_term (Reduction.whd_all env prod) with + let sigma = goalevars evars in + match EConstr.kind sigma (EConstr.of_constr (Reductionops.whd_all env sigma prod)) with | Prod (na, ty, b) -> - if noccurn 1 b then + if noccurn sigma 1 b then let b' = lift (-1) b in let evars, rb = aux evars env b' (pred n) in app_poly env evars pointwise_relation [| ty; b'; rb |] else - let evars, rb = aux evars (Environ.push_rel (LocalAssum (na, ty)) env) b (pred n) in + let evars, rb = aux evars (Environ.push_rel (local_assum (na, ty)) env) b (pred n) in app_poly env evars forall_relation [| ty; mkLambda (na, ty, b); mkLambda (na, ty, rb) |] | _ -> raise Not_found @@ -348,8 +373,10 @@ end) = struct try let evars, found = aux evars env ty (succ (List.length args)) in Some (evars, found, c, ty, arg :: args) with Not_found -> - let ty = whd_all env ty in - find env (mkApp (c, [| arg |])) (Term.prod_applist ty [arg]) args + let sigma = goalevars evars in + let ty = Reductionops.whd_all env sigma ty in + let ty = EConstr.of_constr ty in + find env (mkApp (c, [| arg |])) (prod_applist sigma ty [arg]) args in find env c ty args let unlift_cstr env sigma = function @@ -358,22 +385,21 @@ end) = struct (** Looking up declared rewrite relations (instances of [RewriteRelation]) *) let is_applied_rewrite_relation env sigma rels t = - match kind_of_term t with + match EConstr.kind sigma t with | App (c, args) when Array.length args >= 2 -> - let head = if isApp c then fst (destApp c) else c in - if Globnames.is_global (coq_eq_ref ()) head then None + let head = if isApp sigma c then fst (destApp sigma c) else c in + if Termops.is_global sigma (coq_eq_ref ()) head then None else (try let params, args = Array.chop (Array.length args - 2) args in let env' = Environ.push_rel_context rels env in let sigma = Sigma.Unsafe.of_evar_map sigma in let Sigma ((evar, _), evars, _) = Evarutil.new_type_evar env' sigma Evd.univ_flexible in - let evar = EConstr.Unsafe.to_constr evar in let evars = Sigma.to_evar_map evars in let evars, inst = app_poly env (evars,Evar.Set.empty) rewrite_relation_class [| evar; mkApp (c, params) |] in - let _ = Typeclasses.resolve_one_typeclass env' (goalevars evars) (EConstr.of_constr inst) in + let _ = Typeclasses.resolve_one_typeclass env' (goalevars evars) inst in Some (it_mkProd_or_LetIn t rels) with e when CErrors.noncritical e -> None) | _ -> None @@ -388,7 +414,7 @@ end let type_app_poly env env evd f args = let evars, c = app_poly_nocheck env evd f args in - let evd', t = Typing.type_of env (goalevars evars) (EConstr.of_constr c) in + let evd', t = Typing.type_of env (goalevars evars) c in (evd', cstrevars evars), c module PropGlobal = struct @@ -437,7 +463,7 @@ module TypeGlobal = struct end let sort_of_rel env evm rel = - Reductionops.sort_of_arity env evm (EConstr.of_constr (Retyping.get_type_of env evm (EConstr.of_constr rel))) + Reductionops.sort_of_arity env evm (EConstr.of_constr (Retyping.get_type_of env evm rel)) let is_applied_rewrite_relation = PropGlobal.is_applied_rewrite_relation @@ -460,14 +486,14 @@ let evd_convertible env evd x y = unsolvable constraints remain, so we check that this unification does not introduce any new problem. *) let _, pbs = Evd.extract_all_conv_pbs evd in - let evd' = Evarconv.the_conv_x env (EConstr.of_constr x) (EConstr.of_constr y) evd in + let evd' = Evarconv.the_conv_x env x y evd in let _, pbs' = Evd.extract_all_conv_pbs evd' in if evd' == evd || problem_inclusion pbs' pbs then Some evd' else None with e when CErrors.noncritical e -> None let convertible env evd x y = - Reductionops.is_conv_leq env evd (EConstr.of_constr x) (EConstr.of_constr y) + Reductionops.is_conv_leq env evd x y type hypinfo = { prf : constr; @@ -486,12 +512,14 @@ let error_no_relation () = error "Cannot find a relation to rewrite." let rec decompose_app_rel env evd t = (** Head normalize for compatibility with the old meta mechanism *) - let t = Reductionops.whd_betaiota evd (EConstr.of_constr t) in - match kind_of_term t with + let t = Reductionops.whd_betaiota evd t in + let t = EConstr.of_constr t in + match EConstr.kind evd t with | App (f, [||]) -> assert false | App (f, [|arg|]) -> let (f', argl, argr) = decompose_app_rel env evd arg in - let ty = Typing.unsafe_type_of env evd (EConstr.of_constr argl) in + let ty = Typing.unsafe_type_of env evd argl in + let ty = EConstr.of_constr ty in let f'' = mkLambda (Name default_dependent_ident, ty, mkLambda (Name (Id.of_string "y"), lift 1 ty, mkApp (lift 2 f, [| mkApp (lift 2 f', [| mkRel 2; mkRel 1 |]) |]))) @@ -505,28 +533,29 @@ let rec decompose_app_rel env evd t = let decompose_app_rel env evd t = let (rel, t1, t2) = decompose_app_rel env evd t in - let ty = Retyping.get_type_of env evd (EConstr.of_constr rel) in - let () = if not (Reduction.is_arity env ty) then error_no_relation () in + let ty = Retyping.get_type_of env evd rel in + let ty = EConstr.of_constr ty in + let () = if not (Reductionops.is_arity env evd ty) then error_no_relation () in (rel, t1, t2) let decompose_applied_relation env sigma (c,l) = let open Context.Rel.Declaration in - let ctype = Retyping.get_type_of env sigma (EConstr.of_constr c) in + let ctype = Retyping.get_type_of env sigma c in + let ctype = EConstr.of_constr ctype in let find_rel ty = - let ty = EConstr.of_constr ty in let sigma, cl = Clenv.make_evar_clause env sigma ty in - let l = Miscops.map_bindings EConstr.of_constr l in let sigma = Clenv.solve_evar_clause env sigma true cl l in let { Clenv.cl_holes = holes; Clenv.cl_concl = t } = cl in - let t = EConstr.Unsafe.to_constr t in let (equiv, c1, c2) = decompose_app_rel env sigma t in - let ty1 = Retyping.get_type_of env sigma (EConstr.of_constr c1) in - let ty2 = Retyping.get_type_of env sigma (EConstr.of_constr c2) in + let ty1 = Retyping.get_type_of env sigma c1 in + let ty2 = Retyping.get_type_of env sigma c2 in + let ty1 = EConstr.of_constr ty1 in + let ty2 = EConstr.of_constr ty2 in match evd_convertible env sigma ty1 ty2 with | None -> None | Some sigma -> let sort = sort_of_rel env sigma equiv in - let args = Array.map_of_list (fun h -> EConstr.Unsafe.to_constr h.Clenv.hole_evar) holes in + let args = Array.map_of_list (fun h -> h.Clenv.hole_evar) holes in let value = mkApp (c, args) in Some (sigma, { prf=value; car=ty1; rel = equiv; sort = Sorts.is_prop sort; @@ -535,8 +564,7 @@ let decompose_applied_relation env sigma (c,l) = match find_rel ctype with | Some c -> c | None -> - let ctx,t' = Reductionops.splay_prod env sigma (EConstr.of_constr ctype) in (* Search for underlying eq *) - let t' = EConstr.Unsafe.to_constr t' in + let ctx,t' = Reductionops.splay_prod env sigma ctype in (* Search for underlying eq *) match find_rel (it_mkProd_or_LetIn t' (List.map (fun (n,t) -> local_assum (n, t)) ctx)) with | Some c -> c | None -> error "Cannot find an homogeneous relation to rewrite." @@ -722,16 +750,18 @@ let symmetry env sort rew = let unify_eqn (car, rel, prf, c1, c2, holes, sort) l2r flags env (sigma, cstrs) by t = try let left = if l2r then c1 else c2 in - let sigma = Unification.w_unify ~flags env sigma CONV (EConstr.of_constr left) (EConstr.of_constr t) in + let sigma = Unification.w_unify ~flags env sigma CONV left t in let sigma = Typeclasses.resolve_typeclasses ~filter:(no_constraints cstrs) ~fail:true env sigma in let evd = solve_remaining_by env sigma holes by in - let nf c = Evarutil.nf_evar evd (Reductionops.nf_meta evd c) in + let nf c = nf_evar evd (Reductionops.nf_meta evd c) in let c1 = nf c1 and c2 = nf c2 and rew_car = nf car and rel = nf rel and prf = nf prf in - let ty1 = Retyping.get_type_of env evd (EConstr.of_constr c1) in - let ty2 = Retyping.get_type_of env evd (EConstr.of_constr c2) in + let ty1 = Retyping.get_type_of env evd c1 in + let ty2 = Retyping.get_type_of env evd c2 in + let ty1 = EConstr.of_constr ty1 in + let ty2 = EConstr.of_constr ty2 in let () = if not (convertible env evd ty2 ty1) then raise Reduction.NotConvertible in let rew_evars = evd, cstrs in let rew_prf = RewPrf (rel, prf) in @@ -749,7 +779,7 @@ let unify_abs (car, rel, prf, c1, c2) l2r sort env (sigma, cstrs) t = basically an eq_constr, except when preexisting evars occur in either the lemma or the goal, in which case the eq_constr also solved this evars *) - let sigma = Unification.w_unify ~flags:rewrite_unif_flags env sigma CONV (EConstr.of_constr left) (EConstr.of_constr t) in + let sigma = Unification.w_unify ~flags:rewrite_unif_flags env sigma CONV left t in let rew_evars = sigma, cstrs in let rew_prf = RewPrf (rel, prf) in let rew = { rew_car = car; rew_from = c1; rew_to = c2; rew_prf; rew_evars; } in @@ -766,9 +796,9 @@ let default_flags = { under_lambdas = true; on_morphisms = true; } let get_opt_rew_rel = function RewPrf (rel, prf) -> Some rel | _ -> None let make_eq () = -(*FIXME*) Universes.constr_of_global (Coqlib.build_coq_eq ()) +(*FIXME*) EConstr.of_constr (Universes.constr_of_global (Coqlib.build_coq_eq ())) let make_eq_refl () = -(*FIXME*) Universes.constr_of_global (Coqlib.build_coq_eq_refl ()) +(*FIXME*) EConstr.of_constr (Universes.constr_of_global (Coqlib.build_coq_eq_refl ())) let get_rew_prf r = match r.rew_prf with | RewPrf (rel, prf) -> rel, prf @@ -781,7 +811,7 @@ let poly_subrelation sort = if sort then PropGlobal.subrelation else TypeGlobal.subrelation let resolve_subrelation env avoid car rel sort prf rel' res = - if Termops.eq_constr (fst res.rew_evars) (EConstr.of_constr rel) (EConstr.of_constr rel') then res + if Termops.eq_constr (fst res.rew_evars) rel rel' then res else let evars, app = app_poly_check env res.rew_evars (poly_subrelation sort) [|car; rel; rel'|] in let evars, subrel = new_cstr_evar evars env app in @@ -798,7 +828,8 @@ let resolve_morphism env avoid oldt m ?(fnewt=fun x -> x) args args' (b,cstr) ev let morphargs, morphobjs = Array.chop first args in let morphargs', morphobjs' = Array.chop first args' in let appm = mkApp(m, morphargs) in - let appmtype = Typing.unsafe_type_of env (goalevars evars) (EConstr.of_constr appm) in + let appmtype = Typing.unsafe_type_of env (goalevars evars) appm in + let appmtype = EConstr.of_constr appmtype in let cstrs = List.map (Option.map (fun r -> r.rew_car, get_opt_rew_rel r.rew_prf)) (Array.to_list morphobjs') @@ -818,7 +849,7 @@ let resolve_morphism env avoid oldt m ?(fnewt=fun x -> x) args args' (b,cstr) ev else TypeGlobal.do_subrelation, TypeGlobal.apply_subrelation in Environ.push_named - (LocalDef (Id.of_string "do_subrelation", + (nlocal_def (Id.of_string "do_subrelation", snd (app_poly_sort b env evars dosub [||]), snd (app_poly_nocheck env evars appsub [||]))) env @@ -849,8 +880,8 @@ let resolve_morphism env avoid oldt m ?(fnewt=fun x -> x) args args' (b,cstr) ev x :: acc, x :: subst, evars, sigargs, x :: typeargs') ([], [], evars, sigargs, []) args args' in - let proof = applistc proj (List.rev projargs) in - let newt = applistc m' (List.rev typeargs) in + let proof = applist (proj, List.rev projargs) in + let newt = applist (m', List.rev typeargs) in match respars with [ a, Some r ] -> evars, proof, substl subst a, substl subst r, oldt, fnewt newt | _ -> assert(false) @@ -873,13 +904,13 @@ let apply_rule unify loccs : int pure_strategy = in { strategy = fun { state = occ ; env ; unfresh ; term1 = t ; ty1 = ty ; cstr ; evars } -> - let unif = if isEvar t then None else unify env evars t in + let unif = if isEvar (goalevars evars) t then None else unify env evars t in match unif with | None -> (occ, Fail) | Some rew -> let occ = succ occ in if not (is_occ occ) then (occ, Fail) - else if Termops.eq_constr (fst rew.rew_evars) (EConstr.of_constr t) (EConstr.of_constr rew.rew_to) then (occ, Identity) + else if Termops.eq_constr (fst rew.rew_evars) t rew.rew_to then (occ, Identity) else let res = { rew with rew_car = ty } in let rel, prf = get_rew_prf res in @@ -933,17 +964,18 @@ let reset_env env = Environ.push_rel_context (Environ.rel_context env) env' let fold_match ?(force=false) env sigma c = - let (ci, p, c, brs) = destCase c in - let cty = Retyping.get_type_of env sigma (EConstr.of_constr c) in + let (ci, p, c, brs) = destCase sigma c in + let cty = Retyping.get_type_of env sigma c in + let cty = EConstr.of_constr cty in let dep, pred, exists, (sk,eff) = let env', ctx, body = - let ctx, pred = decompose_lam_assum p in + let ctx, pred = decompose_lam_assum sigma p in let env' = Environ.push_rel_context ctx env in env', ctx, pred in - let sortp = Retyping.get_sort_family_of env' sigma (EConstr.of_constr body) in - let sortc = Retyping.get_sort_family_of env sigma (EConstr.of_constr cty) in - let dep = not (noccurn 1 body) in + let sortp = Retyping.get_sort_family_of env' sigma body in + let sortc = Retyping.get_sort_family_of env sigma cty in + let dep = not (noccurn sigma 1 body) in let pred = if dep then p else it_mkProd_or_LetIn (subst1 mkProp body) (List.tl ctx) in @@ -967,7 +999,8 @@ let fold_match ?(force=false) env sigma c = else raise Not_found in let app = - let ind, args = Inductive.find_rectype env cty in + let ind, args = Inductiveops.find_mrectype env sigma cty in + let args = List.map EConstr.of_constr args in let pars, args = List.chop ci.ci_npar args in let meths = List.map (fun br -> br) (Array.to_list brs) in applist (mkConst sk, pars @ [pred] @ meths @ args @ [c]) @@ -975,10 +1008,11 @@ let fold_match ?(force=false) env sigma c = sk, (if exists then env else reset_env env), app, eff let unfold_match env sigma sk app = - match kind_of_term app with - | App (f', args) when eq_constant (fst (destConst f')) sk -> + match EConstr.kind sigma app with + | App (f', args) when eq_constant (fst (destConst sigma f')) sk -> let v = Environ.constant_value_in (Global.env ()) (sk,Univ.Instance.empty)(*FIXME*) in - Reductionops.whd_beta sigma (EConstr.of_constr (mkApp (v, args))) + let v = EConstr.of_constr v in + EConstr.of_constr (Reductionops.whd_beta sigma (mkApp (v, args))) | _ -> app let is_rew_cast = function RewCast _ -> true | _ -> false @@ -987,7 +1021,7 @@ let subterm all flags (s : 'a pure_strategy) : 'a pure_strategy = let rec aux { state ; env ; unfresh ; term1 = t ; ty1 = ty ; cstr = (prop, cstr) ; evars } = let cstr' = Option.map (fun c -> (ty, Some c)) cstr in - match kind_of_term t with + match EConstr.kind (goalevars evars) t with | App (m, args) -> let rewrite_args state success = let state, (args', evars', progress) = @@ -996,7 +1030,8 @@ let subterm all flags (s : 'a pure_strategy) : 'a pure_strategy = if not (Option.is_empty progress) && not all then state, (None :: acc, evars, progress) else - let argty = Retyping.get_type_of env (goalevars evars) (EConstr.of_constr arg) in + let argty = Retyping.get_type_of env (goalevars evars) arg in + let argty = EConstr.of_constr argty in let state, res = s.strategy { state ; env ; unfresh ; term1 = arg ; ty1 = argty ; @@ -1044,7 +1079,8 @@ let subterm all flags (s : 'a pure_strategy) : 'a pure_strategy = in state, res in if flags.on_morphisms then - let mty = Retyping.get_type_of env (goalevars evars) (EConstr.of_constr m) in + let mty = Retyping.get_type_of env (goalevars evars) m in + let mty = EConstr.of_constr mty in let evars, cstr', m, mty, argsl, args = let argsl = Array.to_list args in let lift = if prop then PropGlobal.lift_cstr else TypeGlobal.lift_cstr in @@ -1071,7 +1107,7 @@ let subterm all flags (s : 'a pure_strategy) : 'a pure_strategy = | x -> x in let res = - { rew_car = Reductionops.hnf_prod_appvect env (goalevars evars) (EConstr.of_constr r.rew_car) (Array.map EConstr.of_constr args); + { rew_car = EConstr.of_constr (Reductionops.hnf_prod_appvect env (goalevars evars) r.rew_car args); rew_from = mkApp(r.rew_from, args); rew_to = mkApp(r.rew_to, args); rew_prf = prf; rew_evars = r.rew_evars } in @@ -1084,10 +1120,12 @@ let subterm all flags (s : 'a pure_strategy) : 'a pure_strategy = in state, res else rewrite_args state None - | Prod (n, x, b) when noccurn 1 b -> + | Prod (n, x, b) when noccurn (goalevars evars) 1 b -> let b = subst1 mkProp b in - let tx = Retyping.get_type_of env (goalevars evars) (EConstr.of_constr x) - and tb = Retyping.get_type_of env (goalevars evars) (EConstr.of_constr b) in + let tx = Retyping.get_type_of env (goalevars evars) x + and tb = Retyping.get_type_of env (goalevars evars) b in + let tx = EConstr.of_constr tx in + let tb = EConstr.of_constr tb in let arr = if prop then PropGlobal.arrow_morphism else TypeGlobal.arrow_morphism in @@ -1097,7 +1135,7 @@ let subterm all flags (s : 'a pure_strategy) : 'a pure_strategy = cstr = (prop,cstr) ; evars = evars' } in let res = match res with - | Success r -> Success { r with rew_to = unfold r.rew_to } + | Success r -> Success { r with rew_to = unfold (goalevars r.rew_evars) r.rew_to } | Fail | Identity -> res in state, res @@ -1118,7 +1156,7 @@ let subterm all flags (s : 'a pure_strategy) : 'a pure_strategy = | Prod (n, dom, codom) -> let lam = mkLambda (n, dom, codom) in let (evars', app), unfold = - if eq_constr (fst evars) (EConstr.of_constr ty) EConstr.mkProp then + if eq_constr (fst evars) ty mkProp then (app_poly_sort prop env evars coq_all [| dom; lam |]), TypeGlobal.unfold_all else let forall = if prop then PropGlobal.coq_forall else TypeGlobal.coq_forall in @@ -1129,7 +1167,7 @@ let subterm all flags (s : 'a pure_strategy) : 'a pure_strategy = cstr = (prop,cstr) ; evars = evars' } in let res = match res with - | Success r -> Success { r with rew_to = unfold r.rew_to } + | Success r -> Success { r with rew_to = unfold (goalevars r.rew_evars) r.rew_to } | Fail | Identity -> res in state, res @@ -1164,8 +1202,9 @@ let subterm all flags (s : 'a pure_strategy) : 'a pure_strategy = | Lambda (n, t, b) when flags.under_lambdas -> let n' = name_app (fun id -> Tactics.fresh_id_in_env unfresh id env) n in let open Context.Rel.Declaration in - let env' = Environ.push_rel (LocalAssum (n', t)) env in - let bty = Retyping.get_type_of env' (goalevars evars) (EConstr.of_constr b) in + let env' = Environ.push_rel (local_assum (n', t)) env in + let bty = Retyping.get_type_of env' (goalevars evars) b in + let bty = EConstr.of_constr bty in let unlift = if prop then PropGlobal.unlift_cstr else TypeGlobal.unlift_cstr in let state, b' = s.strategy { state ; env = env' ; unfresh ; term1 = b ; ty1 = bty ; @@ -1192,7 +1231,8 @@ let subterm all flags (s : 'a pure_strategy) : 'a pure_strategy = in state, res | Case (ci, p, c, brs) -> - let cty = Retyping.get_type_of env (goalevars evars) (EConstr.of_constr c) in + let cty = Retyping.get_type_of env (goalevars evars) c in + let cty = EConstr.of_constr cty in let evars', eqty = app_poly_sort prop env evars coq_eq [| cty |] in let cstr' = Some eqty in let state, c' = s.strategy { state ; env ; unfresh ; @@ -1393,7 +1433,7 @@ module Strategies = let inj_open hint = (); fun sigma -> let ctx = Evd.evar_universe_context_of hint.Autorewrite.rew_ctx in let sigma = Evd.merge_universe_context sigma ctx in - (sigma, (hint.Autorewrite.rew_lemma, NoBindings)) + (sigma, (EConstr.of_constr hint.Autorewrite.rew_lemma, NoBindings)) let old_hints (db : string) : 'a pure_strategy = let rules = Autorewrite.find_rewrites db in @@ -1403,6 +1443,7 @@ module Strategies = let hints (db : string) : 'a pure_strategy = { strategy = fun ({ term1 = t } as input) -> + let t = EConstr.Unsafe.to_constr t in let rules = Autorewrite.find_matches db t in let lemma hint = (inj_open hint, hint.Autorewrite.rew_l2r, hint.Autorewrite.rew_tac) in @@ -1414,9 +1455,10 @@ module Strategies = fun { state = state ; env = env ; term1 = t ; ty1 = ty ; cstr = cstr ; evars = evars } -> let rfn, ckind = Redexpr.reduction_of_red_expr env r in let sigma = Sigma.Unsafe.of_evar_map (goalevars evars) in - let Sigma (t', sigma, _) = rfn.Reductionops.e_redfun env sigma (EConstr.of_constr t) in + let Sigma (t', sigma, _) = rfn.Reductionops.e_redfun env sigma t in + let t' = EConstr.of_constr t' in let evars' = Sigma.to_evar_map sigma in - if Termops.eq_constr evars' (EConstr.of_constr t') (EConstr.of_constr t) then + if Termops.eq_constr evars' t' t then state, Identity else state, Success { rew_car = ty; rew_from = t; rew_to = t'; @@ -1428,14 +1470,16 @@ module Strategies = fun { state ; env ; term1 = t ; ty1 = ty ; cstr ; evars } -> (* let sigma, (c,_) = Tacinterp.interp_open_constr_with_bindings is env (goalevars evars) c in *) let sigma, c = Pretyping.understand_tcc env (goalevars evars) c in + let c = EConstr.of_constr c in let unfolded = - try Tacred.try_red_product env sigma (EConstr.of_constr c) + try Tacred.try_red_product env sigma c with e when CErrors.noncritical e -> error "fold: the term is not unfoldable !" in + let unfolded = EConstr.of_constr unfolded in try - let sigma = Unification.w_unify env sigma CONV ~flags:(Unification.elim_flags ()) (EConstr.of_constr unfolded) (EConstr.of_constr t) in - let c' = Evarutil.nf_evar sigma c in + let sigma = Unification.w_unify env sigma CONV ~flags:(Unification.elim_flags ()) unfolded t in + let c' = nf_evar sigma c in state, Success { rew_car = ty; rew_from = t; rew_to = c'; rew_prf = RewCast DEFAULTcast; rew_evars = (sigma, snd evars) } @@ -1468,7 +1512,8 @@ let rewrite_with l2r flags c occs : strategy = { strategy = } let apply_strategy (s : strategy) env unfresh concl (prop, cstr) evars = - let ty = Retyping.get_type_of env (goalevars evars) (EConstr.of_constr concl) in + let ty = Retyping.get_type_of env (goalevars evars) concl in + let ty = EConstr.of_constr ty in let _, res = s.strategy { state = () ; env ; unfresh ; term1 = concl ; ty1 = ty ; cstr = (prop, Some cstr) ; evars } in @@ -1488,7 +1533,7 @@ type result = (evar_map * constr option * types) option option let cl_rewrite_clause_aux ?(abs=None) strat env avoid sigma concl is_hyp : result = let evdref = ref sigma in - let sort = Typing.e_sort_of env evdref (EConstr.of_constr concl) in + let sort = Typing.e_sort_of env evdref concl in let evars = (!evdref, Evar.Set.empty) in let evars, cstr = let prop, (evars, arrow) = @@ -1508,7 +1553,7 @@ let cl_rewrite_clause_aux ?(abs=None) strat env avoid sigma concl is_hyp : resul | Success res -> let (_, cstrs) = res.rew_evars in let evars' = solve_constraints env res.rew_evars in - let newt = Evarutil.nf_evar evars' res.rew_to in + let newt = EConstr.of_constr (Evarutil.nf_evar evars' (EConstr.Unsafe.to_constr res.rew_to)) in let evars = (* Keep only original evars (potentially instantiated) and goal evars, the rest has been defined and substituted already. *) Evar.Set.fold @@ -1523,13 +1568,14 @@ let cl_rewrite_clause_aux ?(abs=None) strat env avoid sigma concl is_hyp : resul let res = match res.rew_prf with | RewCast c -> None | RewPrf (rel, p) -> - let p = nf_zeta env evars' (EConstr.of_constr (Evarutil.nf_evar evars' p)) in + let p = nf_zeta env evars' (nf_evar evars' p) in + let p = EConstr.of_constr p in let term = match abs with | None -> p | Some (t, ty) -> - let t = Evarutil.nf_evar evars' t in - let ty = Evarutil.nf_evar evars' ty in + let t = nf_evar evars' t in + let ty = nf_evar evars' ty in mkApp (mkLambda (Name (Id.of_string "lemma"), ty, p), [| t |]) in let proof = match is_hyp with @@ -1550,24 +1596,25 @@ let rec insert_dependent env sigma decl accu hyps = match hyps with let assert_replacing id newt tac = let prf = Proofview.Goal.nf_enter { enter = begin fun gl -> let concl = Proofview.Goal.concl gl in + let concl = EConstr.of_constr concl in let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in let ctx = Environ.named_context env in let after, before = List.split_when (NamedDecl.get_id %> Id.equal id) ctx in let nc = match before with | [] -> assert false - | d :: rem -> insert_dependent env sigma (LocalAssum (NamedDecl.get_id d, newt)) [] after @ rem + | d :: rem -> insert_dependent env sigma (nlocal_assum (NamedDecl.get_id d, newt)) [] after @ rem in let env' = Environ.reset_with_named_context (val_of_named_context nc) env in Refine.refine ~unsafe:false { run = begin fun sigma -> let open EConstr in - let Sigma (ev, sigma, p) = Evarutil.new_evar env' sigma (EConstr.of_constr concl) in - let Sigma (ev', sigma, q) = Evarutil.new_evar env sigma (EConstr.of_constr newt) in + let Sigma (ev, sigma, p) = Evarutil.new_evar env' sigma concl in + let Sigma (ev', sigma, q) = Evarutil.new_evar env sigma newt in let map d = let n = NamedDecl.get_id d in - if Id.equal n id then ev' else EConstr.mkVar n + if Id.equal n id then ev' else mkVar n in - let (e, _) = EConstr.destEvar (Sigma.to_evar_map sigma) ev in + let (e, _) = destEvar (Sigma.to_evar_map sigma) ev in Sigma (mkEvar (e, Array.map_of_list map nc), sigma, p +> q) end } end } in @@ -1594,38 +1641,37 @@ let cl_rewrite_clause_newtac ?abs ?origsigma ~progress strat clause = match clause, prf with | Some id, Some p -> let tac = tclTHENLIST [ - Refine.refine ~unsafe:false { run = fun h -> Sigma.here (EConstr.of_constr p) h }; + Refine.refine ~unsafe:false { run = fun h -> Sigma.here p h }; Proofview.Unsafe.tclNEWGOALS gls; ] in Proofview.Unsafe.tclEVARS undef <*> tclTHENFIRST (assert_replacing id newt tac) (beta_hyp id) | Some id, None -> Proofview.Unsafe.tclEVARS undef <*> - convert_hyp_no_check (LocalAssum (id, newt)) <*> + convert_hyp_no_check (nlocal_assum (id, newt)) <*> beta_hyp id | None, Some p -> - let p = EConstr.of_constr p in Proofview.Unsafe.tclEVARS undef <*> Proofview.Goal.enter { enter = begin fun gl -> let env = Proofview.Goal.env gl in let make = { run = begin fun sigma -> - let Sigma (ev, sigma, q) = Evarutil.new_evar env sigma (EConstr.of_constr newt) in - Sigma (EConstr.mkApp (p, [| ev |]), sigma, q) + let Sigma (ev, sigma, q) = Evarutil.new_evar env sigma newt in + Sigma (mkApp (p, [| ev |]), sigma, q) end } in Refine.refine ~unsafe:false make <*> Proofview.Unsafe.tclNEWGOALS gls end } | None, None -> - let newt = EConstr.of_constr newt in Proofview.Unsafe.tclEVARS undef <*> convert_concl_no_check newt DEFAULTcast in Proofview.Goal.nf_enter { enter = begin fun gl -> let concl = Proofview.Goal.concl gl in + let concl = EConstr.of_constr concl in let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in let ty = match clause with | None -> concl - | Some id -> Environ.named_type id env + | Some id -> EConstr.of_constr (Environ.named_type id env) in let env = match clause with | None -> env @@ -1677,6 +1723,7 @@ let cl_rewrite_clause_strat strat clause = let apply_glob_constr c l2r occs = (); fun ({ state = () ; env = env } as input) -> let c sigma = let (sigma, c) = Pretyping.understand_tcc env sigma c in + let c = EConstr.of_constr c in (sigma, (c, NoBindings)) in let flags = general_rewrite_unif_flags () in @@ -1685,6 +1732,7 @@ let apply_glob_constr c l2r occs = (); fun ({ state = () ; env = env } as input) let interp_glob_constr_list env = let make c = (); fun sigma -> let sigma, c = Pretyping.understand_tcc env sigma c in + let c = EConstr.of_constr c in (sigma, (c, NoBindings)) in List.map (fun c -> make c, true, None) @@ -1869,9 +1917,10 @@ let declare_relation ?(binders=[]) a aeq n refl symm trans = let cHole = CHole (Loc.ghost, None, Misctypes.IntroAnonymous, None) -let proper_projection r ty = - let ctx, inst = decompose_prod_assum ty in - let mor, args = destApp inst in +let proper_projection sigma r ty = + let rel_vect n m = Array.init m (fun i -> mkRel(n+m-i)) in + let ctx, inst = decompose_prod_assum sigma ty in + let mor, args = destApp sigma inst in let instarg = mkApp (r, rel_vect 0 (List.length ctx)) in let app = mkApp (Lazy.force PropGlobal.proper_proj, Array.append args [| instarg |]) in @@ -1882,31 +1931,37 @@ let declare_projection n instance_id r = let env = Global.env () in let sigma = Evd.from_env env in let sigma,c = Evd.fresh_global env sigma r in - let ty = Retyping.get_type_of env sigma (EConstr.of_constr c) in - let term = proper_projection c ty in - let sigma, typ = Typing.type_of env sigma (EConstr.of_constr term) in - let ctx, typ = decompose_prod_assum typ in + let c = EConstr.of_constr c in + let ty = Retyping.get_type_of env sigma c in + let ty = EConstr.of_constr ty in + let term = proper_projection sigma c ty in + let sigma, typ = Typing.type_of env sigma term in + let typ = EConstr.of_constr typ in + let ctx, typ = decompose_prod_assum sigma typ in let typ = let n = let rec aux t = - match kind_of_term t with + match EConstr.kind sigma t with | App (f, [| a ; a' ; rel; rel' |]) - when Globnames.is_global (PropGlobal.respectful_ref ()) f -> + when Termops.is_global sigma (PropGlobal.respectful_ref ()) f -> succ (aux rel') | _ -> 0 in let init = - match kind_of_term typ with - App (f, args) when Globnames.is_global (PropGlobal.respectful_ref ()) f -> + match EConstr.kind sigma typ with + App (f, args) when Termops.is_global sigma (PropGlobal.respectful_ref ()) f -> mkApp (f, fst (Array.chop (Array.length args - 2) args)) | _ -> typ in aux init in - let ctx,ccl = Reductionops.splay_prod_n env sigma (3 * n) (EConstr.of_constr typ) + let ctx,ccl = Reductionops.splay_prod_n env sigma (3 * n) typ in + let ccl = EConstr.of_constr ccl in it_mkProd_or_LetIn ccl ctx in let typ = it_mkProd_or_LetIn typ ctx in let pl, ctx = Evd.universe_context sigma in + let typ = EConstr.to_constr sigma typ in + let term = EConstr.to_constr sigma term in let cst = Declare.definition_entry ~types:typ ~poly ~univs:ctx term in @@ -1915,11 +1970,13 @@ let declare_projection n instance_id r = let build_morphism_signature env sigma m = let m,ctx = Constrintern.interp_constr env sigma m in + let m = EConstr.of_constr m in let sigma = Evd.from_ctx ctx in - let t = Typing.unsafe_type_of env sigma (EConstr.of_constr m) in + let t = Typing.unsafe_type_of env sigma m in + let t = EConstr.of_constr t in let cstrs = let rec aux t = - match kind_of_term t with + match EConstr.kind sigma t with | Prod (na, a, b) -> None :: aux b | _ -> [] @@ -1939,21 +1996,21 @@ let build_morphism_signature env sigma m = let morph = e_app_poly env evd PropGlobal.proper_type [| t; sig_; m |] in let evd = solve_constraints env !evd in let evd = Evd.nf_constraints evd in - let m = Evarutil.nf_evars_universes evd morph in + let m = Evarutil.nf_evars_universes evd (EConstr.Unsafe.to_constr morph) in Pretyping.check_evars env Evd.empty evd (EConstr.of_constr m); Evd.evar_universe_context evd, m let default_morphism sign m = let env = Global.env () in let sigma = Evd.from_env env in - let t = Typing.unsafe_type_of env sigma (EConstr.of_constr m) in + let t = Typing.unsafe_type_of env sigma m in + let t = EConstr.of_constr t in let evars, _, sign, cstrs = PropGlobal.build_signature (sigma, Evar.Set.empty) env t (fst sign) (snd sign) in let evars, morph = app_poly_check env evars PropGlobal.proper_type [| t; sign; m |] in - let evars, mor = resolve_one_typeclass env (goalevars evars) (EConstr.of_constr morph) in - let mor = EConstr.Unsafe.to_constr mor in - mor, proper_projection mor morph + let evars, mor = resolve_one_typeclass env (goalevars evars) morph in + mor, proper_projection sigma mor morph let add_setoid global binders a aeq t n = init_setoid (); @@ -2060,23 +2117,22 @@ let unification_rewrite l2r c1 c2 sigma prf car rel but env = in the context *) Unification.w_unify_to_subterm ~flags:rewrite_unif_flags - env sigma (EConstr.of_constr (if l2r then c1 else c2),EConstr.of_constr but) + env sigma ((if l2r then c1 else c2),but) with | ex when Pretype_errors.precatchable_exception ex -> (* ~flags:(true,true) to make Ring work (since it really exploits conversion) *) Unification.w_unify_to_subterm ~flags:rewrite_conv_unif_flags - env sigma (EConstr.of_constr (if l2r then c1 else c2),EConstr.of_constr but) + env sigma ((if l2r then c1 else c2),but) in - let c' = EConstr.Unsafe.to_constr c' in - let nf c = Evarutil.nf_evar sigma c in + let nf c = nf_evar sigma c in let c1 = if l2r then nf c' else nf c1 and c2 = if l2r then nf c2 else nf c' and car = nf car and rel = nf rel in check_evar_map_of_evars_defs sigma; let prf = nf prf in - let prfty = nf (Retyping.get_type_of env sigma (EConstr.of_constr prf)) in + let prfty = nf (EConstr.of_constr (Retyping.get_type_of env sigma prf)) in let sort = sort_of_rel env sigma but in let abs = prf, prfty in let prf = mkRel 1 in @@ -2091,6 +2147,7 @@ let get_hyp gl (c,l) clause l2r = | Some id -> Tacmach.New.pf_get_hyp_typ id gl | None -> Evarutil.nf_evar evars (Tacmach.New.pf_concl gl) in + let but = EConstr.of_constr but in unification_rewrite l2r hi.c1 hi.c2 sigma hi.prf hi.car hi.rel but env let general_rewrite_flags = { under_lambdas = false; on_morphisms = true } @@ -2101,7 +2158,6 @@ let general_rewrite_flags = { under_lambdas = false; on_morphisms = true } (** Setoid rewriting when called with "rewrite" *) let general_s_rewrite cl l2r occs (c,l) ~new_goals = Proofview.Goal.nf_enter { enter = begin fun gl -> - let (c,l) = Miscops.map_with_bindings EConstr.Unsafe.to_constr (c,l) in let abs, evd, res, sort = get_hyp gl (c,l) cl l2r in let unify env evars t = unify_abs res l2r sort env evars t in let app = apply_rule unify occs in @@ -2130,6 +2186,7 @@ let _ = Hook.set Equality.general_setoid_rewrite_clause general_s_rewrite (** [setoid_]{reflexivity,symmetry,transitivity} tactics *) let not_declared env ty rel = + let rel = EConstr.Unsafe.to_constr rel in tclFAIL 0 (str" The relation " ++ Printer.pr_constr_env env Evd.empty rel ++ str" is not a declared " ++ str ty ++ str" relation. Maybe you need to require the Coq.Classes.RelationClasses library") @@ -2139,12 +2196,14 @@ let setoid_proof ty fn fallback = let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in let concl = Proofview.Goal.concl gl in + let concl = EConstr.of_constr concl in Proofview.tclORELSE begin try let rel, _, _ = decompose_app_rel env sigma concl in - let (sigma, t) = Typing.type_of env sigma (EConstr.of_constr rel) in - let car = RelDecl.get_type (List.hd (fst (Reduction.dest_prod env t))) in + let (sigma, t) = Typing.type_of env sigma rel in + let t = EConstr.of_constr t in + let car = snd (List.hd (fst (Reductionops.splay_prod env sigma t))) in (try init_relation_classes () with _ -> raise Not_found); fn env sigma car rel with e -> Proofview.tclZERO e @@ -2180,7 +2239,7 @@ let setoid_reflexivity = tac_open (poly_proof PropGlobal.get_reflexive_proof TypeGlobal.get_reflexive_proof env evm car rel) - (fun c -> tclCOMPLETE (apply (EConstr.of_constr c)))) + (fun c -> tclCOMPLETE (apply c))) (reflexivity_red true) let setoid_symmetry = @@ -2189,7 +2248,7 @@ let setoid_symmetry = tac_open (poly_proof PropGlobal.get_symmetric_proof TypeGlobal.get_symmetric_proof env evm car rel) - (fun c -> apply (EConstr.of_constr c))) + (fun c -> apply c)) (symmetry_red true) let setoid_transitivity c = @@ -2198,15 +2257,17 @@ let setoid_transitivity c = tac_open (poly_proof PropGlobal.get_transitive_proof TypeGlobal.get_transitive_proof env evm car rel) (fun proof -> match c with - | None -> eapply (EConstr.of_constr proof) - | Some c -> apply_with_bindings (EConstr.of_constr proof,ImplicitBindings [ c ]))) + | None -> eapply proof + | Some c -> apply_with_bindings (proof,ImplicitBindings [ c ]))) (transitivity_red true c) let setoid_symmetry_in id = Proofview.V82.tactic (fun gl -> - let ctype = pf_unsafe_type_of gl (EConstr.mkVar id) in - let binders,concl = decompose_prod_assum ctype in - let (equiv, args) = decompose_app concl in + let sigma = project gl in + let ctype = pf_unsafe_type_of gl (mkVar id) in + let ctype = EConstr.of_constr ctype in + let binders,concl = decompose_prod_assum sigma ctype in + let (equiv, args) = decompose_app sigma concl in let rec split_last_two = function | [c1;c2] -> [],(c1, c2) | x::y::z -> let l,res = split_last_two (y::z) in x::l, res @@ -2216,11 +2277,10 @@ let setoid_symmetry_in id = let he,c1,c2 = mkApp (equiv, Array.of_list others),c1,c2 in let new_hyp' = mkApp (he, [| c2 ; c1 |]) in let new_hyp = it_mkProd_or_LetIn new_hyp' binders in - let new_hyp = EConstr.of_constr new_hyp in Proofview.V82.of_tactic (tclTHENLAST (Tactics.assert_after_replacing id new_hyp) - (tclTHENLIST [ intros; setoid_symmetry; apply (EConstr.mkVar id); Tactics.assumption ])) + (tclTHENLIST [ intros; setoid_symmetry; apply (mkVar id); Tactics.assumption ])) gl) let _ = Hook.set Tactics.setoid_reflexivity setoid_reflexivity diff --git a/ltac/rewrite.mli b/ltac/rewrite.mli index bf56eec10..378359d51 100644 --- a/ltac/rewrite.mli +++ b/ltac/rewrite.mli @@ -8,6 +8,7 @@ open Names open Constr +open EConstr open Environ open Constrexpr open Tacexpr @@ -105,13 +106,13 @@ val setoid_symmetry_in : Id.t -> unit Proofview.tactic val setoid_reflexivity : unit Proofview.tactic -val setoid_transitivity : EConstr.constr option -> unit Proofview.tactic +val setoid_transitivity : constr option -> unit Proofview.tactic val apply_strategy : strategy -> Environ.env -> Names.Id.t list -> - Term.constr -> - bool * Term.constr -> + constr -> + bool * constr -> evars -> rewrite_result diff --git a/ltac/tacinterp.ml b/ltac/tacinterp.ml index d2fb2614e..3ed40357e 100644 --- a/ltac/tacinterp.ml +++ b/ltac/tacinterp.ml @@ -1013,6 +1013,7 @@ let interp_constr_with_bindings ist env sigma (c,bl) = let interp_open_constr_with_bindings ist env sigma (c,bl) = let sigma, bl = interp_bindings ist env sigma bl in let sigma, c = interp_open_constr ist env sigma c in + let (c, bl) = Miscops.map_with_bindings EConstr.of_constr (c, bl) in sigma, (c, bl) let loc_of_bindings = function @@ -1027,7 +1028,6 @@ let interp_open_constr_with_bindings_loc ist ((c,_),bl as cb) = let f = { delayed = fun env sigma -> let sigma = Sigma.to_evar_map sigma in let (sigma, c) = interp_open_constr_with_bindings ist env sigma cb in - let c = Miscops.map_with_bindings EConstr.of_constr c in Sigma.Unsafe.of_pair (c, sigma) } in (loc,f) @@ -1899,7 +1899,6 @@ and interp_atomic ist tac : unit Proofview.tactic = let f = { delayed = fun env sigma -> let sigma = Sigma.to_evar_map sigma in let (sigma, c) = interp_open_constr_with_bindings ist env sigma c in - let c = Miscops.map_with_bindings EConstr.of_constr c in Sigma.Unsafe.of_pair (c, sigma) } in (b,m,keep,f)) l in diff --git a/ltac/tacinterp.mli b/ltac/tacinterp.mli index 6f64981ef..2b0324e24 100644 --- a/ltac/tacinterp.mli +++ b/ltac/tacinterp.mli @@ -79,7 +79,7 @@ val interp_bindings : interp_sign -> Environ.env -> Evd.evar_map -> glob_constr_and_expr bindings -> Evd.evar_map * constr bindings val interp_open_constr_with_bindings : interp_sign -> Environ.env -> Evd.evar_map -> - glob_constr_and_expr with_bindings -> Evd.evar_map * constr with_bindings + glob_constr_and_expr with_bindings -> Evd.evar_map * EConstr.constr with_bindings (** Initial call for interpretation *) diff --git a/plugins/decl_mode/decl_proof_instr.ml b/plugins/decl_mode/decl_proof_instr.ml index 54206aa95..5de2c4151 100644 --- a/plugins/decl_mode/decl_proof_instr.ml +++ b/plugins/decl_mode/decl_proof_instr.ml @@ -382,6 +382,9 @@ let enstack_subsubgoals env se stack gls= Array.iteri process gentypes | _ -> () +let nf_meta sigma c = + EConstr.Unsafe.to_constr (Reductionops.nf_meta sigma (EConstr.of_constr c)) + let rec nf_list evd = function [] -> [] @@ -389,7 +392,7 @@ let rec nf_list evd = if meta_defined evd m then nf_list evd others else - (m,Reductionops.nf_meta evd typ)::nf_list evd others + (m,nf_meta evd typ)::nf_list evd others let find_subsubgoal c ctyp skip submetas gls = let env= pf_env gls in @@ -424,7 +427,7 @@ let find_subsubgoal c ctyp skip submetas gls = dfs n end in let nse= try dfs skip with Stack.Empty -> raise Not_found in - nf_list nse.se_evd nse.se_meta_list,Reductionops.nf_meta nse.se_evd (mkMeta 0) + nf_list nse.se_evd nse.se_meta_list,nf_meta nse.se_evd (mkMeta 0) let concl_refiner metas body gls = let concl = pf_concl gls in diff --git a/plugins/setoid_ring/newring.ml b/plugins/setoid_ring/newring.ml index 089e76d7a..4492b854b 100644 --- a/plugins/setoid_ring/newring.ml +++ b/plugins/setoid_ring/newring.ml @@ -416,11 +416,18 @@ let theory_to_obj : ring_info -> obj = let setoid_of_relation env evd a r = + let a = EConstr.of_constr a in + let r = EConstr.of_constr r in try let evm = !evd in let evm, refl = Rewrite.get_reflexive_proof env evm a r in let evm, sym = Rewrite.get_symmetric_proof env evm a r in let evm, trans = Rewrite.get_transitive_proof env evm a r in + let refl = EConstr.Unsafe.to_constr refl in + let sym = EConstr.Unsafe.to_constr sym in + let trans = EConstr.Unsafe.to_constr trans in + let a = EConstr.Unsafe.to_constr a in + let r = EConstr.Unsafe.to_constr r in evd := evm; lapp coq_mk_Setoid [|a ; r ; refl; sym; trans |] with Not_found -> @@ -498,22 +505,32 @@ let ring_equality env evd (r,add,mul,opp,req) = (setoid,op_morph) | _ -> let setoid = setoid_of_relation (Global.env ()) evd r req in - let signature = [Some (r,Some req);Some (r,Some req)],Some(r,Some req) in + let signature = + let r = EConstr.of_constr r in + let req = EConstr.of_constr req in + [Some (r,Some req);Some (r,Some req)],Some(r,Some req) + in +(* let signature = [Some (r,Some req);Some (r,Some req)],Some(r,Some req) in *) let add_m, add_m_lem = - try Rewrite.default_morphism signature add + try Rewrite.default_morphism signature (EConstr.of_constr add) with Not_found -> error "ring addition should be declared as a morphism" in + let add_m_lem = EConstr.Unsafe.to_constr add_m_lem in let mul_m, mul_m_lem = - try Rewrite.default_morphism signature mul + try Rewrite.default_morphism signature (EConstr.of_constr mul) with Not_found -> error "ring multiplication should be declared as a morphism" in + let mul_m_lem = EConstr.Unsafe.to_constr mul_m_lem in let op_morph = match opp with | Some opp -> (let opp_m,opp_m_lem = - try Rewrite.default_morphism ([Some(r,Some req)],Some(r,Some req)) opp + let r = EConstr.of_constr r in + let req = EConstr.of_constr req in + try Rewrite.default_morphism ([Some(r,Some req)],Some(r,Some req)) (EConstr.of_constr opp) with Not_found -> error "ring opposite should be declared as a morphism" in + let opp_m_lem = EConstr.Unsafe.to_constr opp_m_lem in let op_morph = op_morph r add mul opp req add_m_lem mul_m_lem opp_m_lem in Flags.if_verbose @@ -895,11 +912,15 @@ let field_equality evd r inv req = mkApp(Universes.constr_of_global (Coqlib.build_coq_eq_data()).congr,[|r;r;inv|]) | _ -> let _setoid = setoid_of_relation (Global.env ()) evd r req in + let r = EConstr.of_constr r in + let req = EConstr.of_constr req in let signature = [Some (r,Some req)],Some(r,Some req) in + let inv = EConstr.of_constr inv in let inv_m, inv_m_lem = try Rewrite.default_morphism signature inv with Not_found -> error "field inverse should be declared as a morphism" in + let inv_m_lem = EConstr.Unsafe.to_constr inv_m_lem in inv_m_lem let add_field_theory name (sigma,fth) eqth morphth cst_tac inj (pre,post) power sign odiv = diff --git a/pretyping/reductionops.ml b/pretyping/reductionops.ml index 480ec2319..31354217f 100644 --- a/pretyping/reductionops.ml +++ b/pretyping/reductionops.ml @@ -1592,8 +1592,9 @@ let meta_instance sigma b = EConstr.of_constr (instance sigma c_sigma b.rebus) let nf_meta sigma c = + let c = EConstr.Unsafe.to_constr c in let cl = mk_freelisted c in - EConstr.Unsafe.to_constr (meta_instance sigma { cl with rebus = EConstr.of_constr cl.rebus }) + meta_instance sigma { cl with rebus = EConstr.of_constr cl.rebus } (* Instantiate metas that create beta/iota redexes *) diff --git a/pretyping/reductionops.mli b/pretyping/reductionops.mli index e67fad3fd..1e6527b29 100644 --- a/pretyping/reductionops.mli +++ b/pretyping/reductionops.mli @@ -296,5 +296,5 @@ val whd_betaiota_deltazeta_for_iota_state : (** {6 Meta-related reduction functions } *) val meta_instance : evar_map -> EConstr.constr freelisted -> EConstr.constr -val nf_meta : evar_map -> constr -> constr +val nf_meta : evar_map -> EConstr.constr -> EConstr.constr val meta_reducible_instance : evar_map -> EConstr.constr freelisted -> EConstr.constr diff --git a/pretyping/unification.ml b/pretyping/unification.ml index bc59a4108..81d9ecad5 100644 --- a/pretyping/unification.ml +++ b/pretyping/unification.ml @@ -1272,13 +1272,15 @@ let w_coerce env evd mv c = let unify_to_type env sigma flags c status u = let sigma, c = refresh_universes (Some false) env sigma c in - let t = get_type_of env sigma (EConstr.of_constr (nf_meta sigma c)) in - let t = nf_betaiota sigma (EConstr.of_constr (nf_meta sigma t)) in - unify_0 env sigma CUMUL flags (EConstr.of_constr t) (EConstr.of_constr u) + let c = EConstr.of_constr c in + let t = get_type_of env sigma (nf_meta sigma c) in + let t = EConstr.of_constr t in + let t = nf_betaiota sigma (nf_meta sigma t) in + let t = EConstr.of_constr t in + unify_0 env sigma CUMUL flags t u let unify_type env sigma flags mv status c = let mvty = Typing.meta_type sigma mv in - let mvty = EConstr.Unsafe.to_constr mvty in let mvty = nf_meta sigma mvty in unify_to_type env sigma (set_flags_for_type flags) diff --git a/proofs/clenv.ml b/proofs/clenv.ml index 514fc27e8..d98669660 100644 --- a/proofs/clenv.ml +++ b/proofs/clenv.ml @@ -44,7 +44,7 @@ type clausenv = { let cl_env ce = ce.env let cl_sigma ce = ce.evd -let clenv_nf_meta clenv c = EConstr.of_constr (nf_meta clenv.evd (EConstr.Unsafe.to_constr c)) +let clenv_nf_meta clenv c = nf_meta clenv.evd c let clenv_term clenv c = meta_instance clenv.evd c let clenv_meta_type clenv mv = Typing.meta_type clenv.evd mv let clenv_value clenv = meta_instance clenv.evd clenv.templval -- cgit v1.2.3 From 2cd0648e003308a000f9f89c898bce4d15fc94a1 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sun, 20 Nov 2016 20:01:22 +0100 Subject: Tauto API using EConstr. --- ltac/tauto.ml | 37 +++++++++++++++---------------------- 1 file changed, 15 insertions(+), 22 deletions(-) diff --git a/ltac/tauto.ml b/ltac/tauto.ml index e3f5342de..a4faf438a 100644 --- a/ltac/tauto.ml +++ b/ltac/tauto.ml @@ -7,6 +7,7 @@ (************************************************************************) open Term +open EConstr open Hipattern open Names open Pp @@ -24,7 +25,7 @@ let () = Mltop.add_known_module tauto_plugin let assoc_var s ist = let v = Id.Map.find (Names.Id.of_string s) ist.lfun in match Value.to_constr v with - | Some c -> c + | Some c -> EConstr.of_constr c | None -> failwith "tauto: anomaly" (** Parametrization of tauto *) @@ -113,7 +114,7 @@ let split = Tactics.split_with_bindings false [Misctypes.NoBindings] let is_empty _ ist = Proofview.tclEVARMAP >>= fun sigma -> - if is_empty_type sigma (EConstr.of_constr (assoc_var "X1" ist)) then idtac else fail + if is_empty_type sigma (assoc_var "X1" ist) then idtac else fail (* Strictly speaking, this exceeds the propositional fragment as it matches also equality types (and solves them if a reflexivity) *) @@ -121,13 +122,12 @@ let is_unit_or_eq _ ist = Proofview.tclEVARMAP >>= fun sigma -> let flags = assoc_flags ist in let test = if flags.strict_unit then is_unit_type else is_unit_or_eq_type in - if test sigma (EConstr.of_constr (assoc_var "X1" ist)) then idtac else fail + if test sigma (assoc_var "X1" ist) then idtac else fail -let bugged_is_binary t = - let t = EConstr.Unsafe.to_constr t in - isApp t && - let (hdapp,args) = decompose_app t in - match (kind_of_term hdapp) with +let bugged_is_binary sigma t = + isApp sigma t && + let (hdapp,args) = decompose_app sigma t in + match EConstr.kind sigma hdapp with | Ind (ind,u) -> let (mib,mip) = Global.lookup_inductive ind in Int.equal mib.Declarations.mind_nparams 2 @@ -139,8 +139,7 @@ let is_conj _ ist = Proofview.tclEVARMAP >>= fun sigma -> let flags = assoc_flags ist in let ind = assoc_var "X1" ist in - let ind = EConstr.of_constr ind in - if (not flags.binary_mode_bugged_detection || bugged_is_binary ind) && + if (not flags.binary_mode_bugged_detection || bugged_is_binary sigma ind) && is_conjunction sigma ~strict:flags.strict_in_hyp_and_ccl ~onlybinary:flags.binary_mode ind @@ -151,7 +150,6 @@ let flatten_contravariant_conj _ ist = Proofview.tclEVARMAP >>= fun sigma -> let flags = assoc_flags ist in let typ = assoc_var "X1" ist in - let typ = EConstr.of_constr typ in let c = assoc_var "X2" ist in let hyp = assoc_var "id" ist in match match_with_conjunction sigma @@ -159,12 +157,10 @@ let flatten_contravariant_conj _ ist = ~onlybinary:flags.binary_mode typ with | Some (_,args) -> - let args = List.map EConstr.Unsafe.to_constr args in let newtyp = List.fold_right mkArrow args c in - let newtyp = EConstr.of_constr newtyp in let intros = tclMAP (fun _ -> intro) args in - let by = tclTHENLIST [intros; apply (EConstr.of_constr hyp); split; assumption] in - tclTHENLIST [assert_ ~by newtyp; clear (destVar hyp)] + let by = tclTHENLIST [intros; apply hyp; split; assumption] in + tclTHENLIST [assert_ ~by newtyp; clear (destVar sigma hyp)] | _ -> fail (** Dealing with disjunction *) @@ -173,8 +169,7 @@ let is_disj _ ist = Proofview.tclEVARMAP >>= fun sigma -> let flags = assoc_flags ist in let t = assoc_var "X1" ist in - let t = EConstr.of_constr t in - if (not flags.binary_mode_bugged_detection || bugged_is_binary t) && + if (not flags.binary_mode_bugged_detection || bugged_is_binary sigma t) && is_disjunction sigma ~strict:flags.strict_in_hyp_and_ccl ~onlybinary:flags.binary_mode t @@ -185,9 +180,7 @@ let flatten_contravariant_disj _ ist = Proofview.tclEVARMAP >>= fun sigma -> let flags = assoc_flags ist in let typ = assoc_var "X1" ist in - let typ = EConstr.of_constr typ in let c = assoc_var "X2" ist in - let c = EConstr.of_constr c in let hyp = assoc_var "id" ist in match match_with_disjunction sigma ~strict:flags.strict_in_contravariant_hyp @@ -195,13 +188,13 @@ let flatten_contravariant_disj _ ist = typ with | Some (_,args) -> let map i arg = - let typ = EConstr.mkArrow arg c in + let typ = mkArrow arg c in let ci = Tactics.constructor_tac false None (succ i) Misctypes.NoBindings in - let by = tclTHENLIST [intro; apply (EConstr.of_constr hyp); ci; assumption] in + let by = tclTHENLIST [intro; apply hyp; ci; assumption] in assert_ ~by typ in let tacs = List.mapi map args in - let tac0 = clear (destVar hyp) in + let tac0 = clear (destVar sigma hyp) in tclTHEN (tclTHENLIST tacs) tac0 | _ -> fail -- cgit v1.2.3 From d4b344acb23f19b089098b7788f37ea22bc07b81 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sun, 20 Nov 2016 20:09:26 +0100 Subject: Eliminating parts of the right-hand side compatibility layer --- engine/termops.ml | 15 +++++++++------ engine/termops.mli | 12 ++++++------ ltac/rewrite.ml | 5 ++--- plugins/btauto/refl_btauto.ml | 2 +- plugins/cc/cctac.ml | 1 + plugins/decl_mode/decl_proof_instr.ml | 5 ++++- plugins/extraction/extraction.ml | 4 ++-- plugins/firstorder/unify.ml | 3 ++- plugins/fourier/fourierR.ml | 1 + plugins/quote/quote.ml | 2 +- pretyping/cases.ml | 2 +- pretyping/coercion.ml | 4 ++-- pretyping/constr_matching.ml | 2 +- pretyping/detyping.ml | 3 ++- pretyping/evarconv.ml | 6 +++--- pretyping/evarsolve.ml | 14 ++++++-------- pretyping/inductiveops.ml | 4 ++-- pretyping/inductiveops.mli | 6 +++--- pretyping/pretyping.ml | 3 +-- pretyping/reductionops.ml | 14 +++++++------- pretyping/reductionops.mli | 2 +- pretyping/retyping.ml | 11 +++++------ pretyping/tacred.ml | 14 ++++---------- pretyping/typeclasses.ml | 8 ++++---- pretyping/typing.ml | 3 +-- pretyping/unification.ml | 13 ++++--------- proofs/clenv.ml | 4 ++-- proofs/clenvtac.ml | 2 +- tactics/class_tactics.ml | 1 - tactics/equality.ml | 13 +++---------- tactics/hints.ml | 5 ----- tactics/inv.ml | 1 - tactics/tactics.ml | 6 ++---- toplevel/class.ml | 9 +++++---- toplevel/obligations.ml | 2 +- toplevel/search.ml | 28 ++++++++++++++-------------- toplevel/vernacentries.ml | 3 ++- 37 files changed, 106 insertions(+), 127 deletions(-) diff --git a/engine/termops.ml b/engine/termops.ml index ecc6ab68e..abaa409bd 100644 --- a/engine/termops.ml +++ b/engine/termops.ml @@ -276,10 +276,11 @@ let last_arg sigma c = match EConstr.kind sigma c with (* Get the last arg of an application *) let decompose_app_vect sigma c = match EConstr.kind sigma c with - | App (f,cl) -> (EConstr.Unsafe.to_constr f, Array.map EConstr.Unsafe.to_constr cl) - | _ -> (EConstr.Unsafe.to_constr c,[||]) + | 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 @@ -698,13 +699,13 @@ let rec subst_meta bl c = let rec strip_outer_cast sigma c = match EConstr.kind sigma c with | Cast (c,_,_) -> strip_outer_cast sigma c - | _ -> EConstr.Unsafe.to_constr 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 (EConstr.of_constr (strip_outer_cast sigma f)) with + 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 @@ -744,15 +745,17 @@ let my_prefix_application sigma eq_fun (k,c) by_c t = works if [c] has rels *) 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 sigma eq_fun kc t with | Some x -> x | None -> if eq_fun sigma c t then EConstr.mkRel k else - EConstr.map_with_binders sigma (fun (k,c) -> (k+1, EConstr.Vars.lift 1 c)) substrec kc t + EConstr.map_with_binders sigma (fun (k,c) -> (k+1,lift 1 c)) substrec kc t in - EConstr.Unsafe.to_constr (substrec (1,c) t) + substrec (1,c) t let subst_term sigma c t = subst_term_gen sigma EConstr.eq_constr c t diff --git a/engine/termops.mli b/engine/termops.mli index 05604beda..426b5f34d 100644 --- a/engine/termops.mli +++ b/engine/termops.mli @@ -144,7 +144,7 @@ val pop : EConstr.t -> constr (** [subst_term_gen eq d c] replaces [d] by [Rel 1] in [c] using [eq] as equality *) val subst_term_gen : Evd.evar_map -> - (Evd.evar_map -> EConstr.t -> EConstr.t -> bool) -> EConstr.t -> EConstr.t -> constr + (Evd.evar_map -> EConstr.t -> EConstr.t -> bool) -> EConstr.t -> EConstr.t -> EConstr.constr (** [replace_term_gen eq d e c] replaces [d] by [e] in [c] using [eq] as equality *) @@ -153,7 +153,7 @@ val replace_term_gen : EConstr.t -> EConstr.t -> EConstr.t -> constr (** [subst_term d c] replaces [d] by [Rel 1] in [c] *) -val subst_term : Evd.evar_map -> EConstr.t -> EConstr.t -> constr +val subst_term : Evd.evar_map -> EConstr.t -> EConstr.t -> EConstr.constr (** [replace_term d e c] replaces [d] by [e] in [c] *) val replace_term : Evd.evar_map -> EConstr.t -> EConstr.t -> EConstr.t -> constr @@ -174,7 +174,7 @@ val prod_applist : Evd.evar_map -> EConstr.t -> EConstr.t list -> EConstr.t (** 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 -> EConstr.t -> constr +val strip_outer_cast : Evd.evar_map -> EConstr.t -> EConstr.constr exception CannotFilter @@ -204,10 +204,10 @@ val nb_prod_modulo_zeta : Evd.evar_map -> EConstr.t -> int val last_arg : Evd.evar_map -> EConstr.t -> EConstr.constr (** Force the decomposition of a term as an applicative one *) -val decompose_app_vect : Evd.evar_map -> EConstr.t -> constr * constr array +val decompose_app_vect : Evd.evar_map -> EConstr.t -> EConstr.constr * EConstr.constr array -val adjust_app_list_size : constr -> constr list -> constr -> constr list -> - (constr * constr list * constr * constr list) +val adjust_app_list_size : EConstr.constr -> EConstr.constr list -> EConstr.constr -> EConstr.constr list -> + (EConstr.constr * EConstr.constr list * EConstr.constr * EConstr.constr list) val adjust_app_array_size : EConstr.constr -> EConstr.constr array -> EConstr.constr -> EConstr.constr array -> (EConstr.constr * EConstr.constr array * EConstr.constr * EConstr.constr array) diff --git a/ltac/rewrite.ml b/ltac/rewrite.ml index ecb653587..119e872f8 100644 --- a/ltac/rewrite.ml +++ b/ltac/rewrite.ml @@ -323,7 +323,7 @@ end) = struct | App (f, [| a; b; relb |]) when Termops.is_global sigma (pointwise_relation_ref ()) f -> decomp_pointwise sigma (pred n) relb | App (f, [| a; b; arelb |]) when Termops.is_global sigma (forall_relation_ref ()) f -> - decomp_pointwise sigma (pred n) (EConstr.of_constr (Reductionops.beta_applist sigma (arelb, [mkRel 1]))) + decomp_pointwise sigma (pred n) (Reductionops.beta_applist sigma (arelb, [mkRel 1])) | _ -> invalid_arg "decomp_pointwise" let rec apply_pointwise sigma rel = function @@ -332,7 +332,7 @@ end) = struct | App (f, [| a; b; relb |]) when Termops.is_global sigma (pointwise_relation_ref ()) f -> apply_pointwise sigma relb args | App (f, [| a; b; arelb |]) when Termops.is_global sigma (forall_relation_ref ()) f -> - apply_pointwise sigma (EConstr.of_constr (Reductionops.beta_applist sigma (arelb, [arg]))) args + apply_pointwise sigma (Reductionops.beta_applist sigma (arelb, [arg])) args | _ -> invalid_arg "apply_pointwise") | [] -> rel @@ -1000,7 +1000,6 @@ let fold_match ?(force=false) env sigma c = in let app = let ind, args = Inductiveops.find_mrectype env sigma cty in - let args = List.map EConstr.of_constr args in let pars, args = List.chop ci.ci_npar args in let meths = List.map (fun br -> br) (Array.to_list brs) in applist (mkConst sk, pars @ [pred] @ meths @ args @ [c]) diff --git a/plugins/btauto/refl_btauto.ml b/plugins/btauto/refl_btauto.ml index 27398cf65..4d2859fb0 100644 --- a/plugins/btauto/refl_btauto.ml +++ b/plugins/btauto/refl_btauto.ml @@ -15,7 +15,7 @@ let get_inductive dir s = Lazy.from_fun (fun () -> Globnames.destIndRef (glob_ref ())) let decomp_term sigma (c : Term.constr) = - Term.kind_of_term (Termops.strip_outer_cast sigma (EConstr.of_constr c)) + Term.kind_of_term (EConstr.Unsafe.to_constr (Termops.strip_outer_cast sigma (EConstr.of_constr c))) let lapp c v = Term.mkApp (Lazy.force c, v) diff --git a/plugins/cc/cctac.ml b/plugins/cc/cctac.ml index a12ef00ec..6295e004e 100644 --- a/plugins/cc/cctac.ml +++ b/plugins/cc/cctac.ml @@ -87,6 +87,7 @@ let rec decompose_term env sigma t= (Appli (Symb (mkConst (Projection.constant p')), decompose_term env sigma c)) | _ -> let t = Termops.strip_outer_cast sigma (EConstr.of_constr t) in + let t = EConstr.Unsafe.to_constr t in if closed0 t then Symb t else raise Not_found (* decompose equality in members and type *) diff --git a/plugins/decl_mode/decl_proof_instr.ml b/plugins/decl_mode/decl_proof_instr.ml index 5de2c4151..e73166be2 100644 --- a/plugins/decl_mode/decl_proof_instr.ml +++ b/plugins/decl_mode/decl_proof_instr.ml @@ -786,7 +786,7 @@ let rec consider_match may_intro introduced available expected gls = let consider_tac c hyps gls = let c = EConstr.of_constr c in - match kind_of_term (strip_outer_cast (project gls) c) with + match EConstr.kind (project gls) (strip_outer_cast (project gls) c) with Var id -> consider_match false [] [id] hyps gls | _ -> @@ -811,6 +811,9 @@ let rec take_tac wits gls = (* tactics for define *) +let subst_term sigma c t = + EConstr.Unsafe.to_constr (subst_term sigma c t) + let rec build_function sigma args body = match args with st::rest -> diff --git a/plugins/extraction/extraction.ml b/plugins/extraction/extraction.ml index e9fd40722..460157130 100644 --- a/plugins/extraction/extraction.ml +++ b/plugins/extraction/extraction.ml @@ -42,11 +42,11 @@ let none = Evd.empty let type_of env c = let polyprop = (lang() == Haskell) in - Retyping.get_type_of ~polyprop env none (EConstr.of_constr (strip_outer_cast none (EConstr.of_constr c))) + Retyping.get_type_of ~polyprop env none (strip_outer_cast none (EConstr.of_constr c)) let sort_of env c = let polyprop = (lang() == Haskell) in - Retyping.get_sort_family_of ~polyprop env none (EConstr.of_constr (strip_outer_cast none (EConstr.of_constr c))) + Retyping.get_sort_family_of ~polyprop env none (strip_outer_cast none (EConstr.of_constr c)) (*S Generation of flags and signatures. *) diff --git a/plugins/firstorder/unify.ml b/plugins/firstorder/unify.ml index fb237f29b..205cb282d 100644 --- a/plugins/firstorder/unify.ml +++ b/plugins/firstorder/unify.ml @@ -21,7 +21,8 @@ exception UFAIL of constr*constr to the equation set. Raises UFAIL with a pair of terms *) -let strip_outer_cast t = strip_outer_cast Evd.empty (EConstr.of_constr t) (** FIXME *) +let strip_outer_cast t = + EConstr.Unsafe.to_constr (strip_outer_cast Evd.empty (EConstr.of_constr t)) (** FIXME *) let unif t1 t2= let evd = Evd.empty in (** FIXME *) diff --git a/plugins/fourier/fourierR.ml b/plugins/fourier/fourierR.ml index dbb5cc2de..f9802ee5e 100644 --- a/plugins/fourier/fourierR.ml +++ b/plugins/fourier/fourierR.ml @@ -471,6 +471,7 @@ let rec fourier () = let sigma = Tacmach.New.project gl in Coqlib.check_required_library ["Coq";"fourier";"Fourier"]; let goal = Termops.strip_outer_cast sigma (EConstr.of_constr concl) in + let goal = EConstr.Unsafe.to_constr goal in let fhyp=Id.of_string "new_hyp_for_fourier" in (* si le but est une inéquation, on introduit son contraire, et le but à prouver devient False *) diff --git a/plugins/quote/quote.ml b/plugins/quote/quote.ml index 09e77598a..04a747fb4 100644 --- a/plugins/quote/quote.ml +++ b/plugins/quote/quote.ml @@ -183,7 +183,7 @@ type inversion_scheme = { let i_can't_do_that () = error "Quote: not a simple fixpoint" -let decomp_term gl c = kind_of_term (Termops.strip_outer_cast (Tacmach.New.project gl) (EConstr.of_constr c)) +let decomp_term gl c = kind_of_term (EConstr.Unsafe.to_constr (Termops.strip_outer_cast (Tacmach.New.project gl) (EConstr.of_constr c))) (*s [compute_lhs typ i nargsi] builds the term \texttt{(C ?nargsi ... ?2 ?1)}, where \texttt{C} is the [i]-th constructor of inductive diff --git a/pretyping/cases.ml b/pretyping/cases.ml index 119e92c82..76ced2b1d 100644 --- a/pretyping/cases.ml +++ b/pretyping/cases.ml @@ -2335,7 +2335,7 @@ let abstract_tomatch env sigma tomatchs tycon = Rel n -> (lift lenctx c, lift_tomatch_type lenctx t) :: prev, ctx, names, tycon | _ -> let tycon = Option.map - (fun t -> EConstr.of_constr (subst_term sigma (lift 1 c) (lift 1 t))) tycon in + (fun t -> subst_term sigma (lift 1 c) (lift 1 t)) tycon in let name = next_ident_away (Id.of_string "filtered_var") names in (mkRel 1, lift_tomatch_type (succ lenctx) t) :: lift_ctx 1 prev, local_def (Name name, lift lenctx c, lift lenctx $ type_of_tomatch t) :: ctx, diff --git a/pretyping/coercion.ml b/pretyping/coercion.ml index d67976232..48f7be4bb 100644 --- a/pretyping/coercion.ml +++ b/pretyping/coercion.ml @@ -502,8 +502,8 @@ let rec inh_conv_coerce_to_fail loc env evd rigidonly v t c1 = let v2 = Option.map (fun v -> beta_applist evd' (lift 1 v,[v1])) v in let t2 = match v2 with | None -> subst_term evd' v1 t2 - | Some v2 -> Retyping.get_type_of env1 evd' (EConstr.of_constr v2) in - let (evd'',v2') = inh_conv_coerce_to_fail loc env1 evd' rigidonly (Option.map EConstr.of_constr v2) (EConstr.of_constr t2) u2 in + | Some v2 -> EConstr.of_constr (Retyping.get_type_of env1 evd' v2) in + let (evd'',v2') = inh_conv_coerce_to_fail loc env1 evd' rigidonly v2 t2 u2 in (evd'', Option.map (fun v2' -> mkLambda (name, u1, v2')) v2') | _ -> raise (NoCoercionNoUnifier (best_failed_evd,e)) diff --git a/pretyping/constr_matching.ml b/pretyping/constr_matching.ml index 06daa5116..55612aa66 100644 --- a/pretyping/constr_matching.ml +++ b/pretyping/constr_matching.ml @@ -185,7 +185,7 @@ let matches_core env sigma convert allow_partial_app allow_bound_rels else false) in let rec sorec ctx env subst p t = - let cT = EConstr.of_constr (strip_outer_cast sigma t) in + let cT = strip_outer_cast sigma t in match p, EConstr.kind sigma cT with | PSoApp (n,args),m -> let fold (ans, seen) = function diff --git a/pretyping/detyping.ml b/pretyping/detyping.ml index c0611dcec..87561868f 100644 --- a/pretyping/detyping.ml +++ b/pretyping/detyping.ml @@ -267,7 +267,7 @@ let rec decomp_branch tags nal b (avoid,env as e) sigma c = | [] -> (List.rev nal,(e,c)) | b::tags -> let na,c,f,body,t = - match kind_of_term (strip_outer_cast sigma (EConstr.of_constr c)), b with + match kind_of_term (EConstr.Unsafe.to_constr (strip_outer_cast sigma (EConstr.of_constr c))), b with | Lambda (na,t,c),false -> na,c,compute_displayed_let_name_in,None,Some t | LetIn (na,b,t,c),true -> na,c,compute_displayed_name_in,Some b,Some t @@ -503,6 +503,7 @@ let rec detype flags avoid env sigma t = let body = pb.Declarations.proj_body in let ty = Retyping.get_type_of (snd env) sigma (EConstr.of_constr c) in let ((ind,u), args) = Inductiveops.find_mrectype (snd env) sigma (EConstr.of_constr ty) in + let args = List.map EConstr.Unsafe.to_constr args in let body' = strip_lam_assum body in let body' = subst_instance_constr u body' in substl (c :: List.rev args) body' diff --git a/pretyping/evarconv.ml b/pretyping/evarconv.ml index a968af7ff..9675ae2ea 100644 --- a/pretyping/evarconv.ml +++ b/pretyping/evarconv.ml @@ -170,7 +170,7 @@ let check_conv_record env sigma (t1,sk1) (t2,sk2) = let (i,u), ind_args = try Inductiveops.find_mrectype env sigma (EConstr.of_constr ty) with _ -> raise Not_found - in Stack.append_app_list (List.map EConstr.of_constr ind_args) Stack.empty, c, sk1 + in Stack.append_app_list ind_args Stack.empty, c, sk1 | None -> match Stack.strip_n_app nparams sk1 with | Some (params1, c1, extra_args1) -> params1, c1, extra_args1 @@ -188,7 +188,7 @@ let check_conv_record env sigma (t1,sk1) (t2,sk2) = let t' = subst_univs_level_constr subst t' in let bs' = List.map (EConstr.of_constr %> subst_univs_level_constr subst) bs in let h, _ = decompose_app_vect sigma t' in - ctx',(EConstr.of_constr h, t2),c',bs',(Stack.append_app_list params Stack.empty,params1), + ctx',(h, t2),c',bs',(Stack.append_app_list params Stack.empty,params1), (Stack.append_app_list us Stack.empty,us2),(extra_args1,extra_args2),c1, (n, Stack.zip sigma (t2,sk2)) @@ -907,7 +907,7 @@ and conv_record trs env evd (ctx,(h,h2),c,bs,(params,params1),(us,us2),(sk1,sk2) (fun i -> exact_ise_stack2 env i (evar_conv_x trs) sk1 sk2); test; (fun i -> evar_conv_x trs env i CONV h2 - (EConstr.of_constr (fst (decompose_app_vect i (substl ks h)))))] + (fst (decompose_app_vect i (substl ks h))))] else UnifFailure(evd,(*dummy*)NotSameHead) and eta_constructor ts env evd sk1 ((ind, i), u) sk2 term2 = diff --git a/pretyping/evarsolve.ml b/pretyping/evarsolve.ml index 772571926..65b6d287d 100644 --- a/pretyping/evarsolve.ml +++ b/pretyping/evarsolve.ml @@ -518,7 +518,6 @@ let is_unification_pattern (env,nb) evd f l t = let solve_pattern_eqn env sigma l c = let c' = List.fold_right (fun a c -> let c' = subst_term sigma (lift 1 a) (lift 1 c) in - let c' = EConstr.of_constr c' in match EConstr.kind sigma a with (* Rem: if [a] links to a let-in, do as if it were an assumption *) | Rel n -> @@ -557,7 +556,7 @@ let make_projectable_subst aliases sigma evi args = | LocalAssum (id,c), a::rest -> let cstrs = let a',args = decompose_app_vect sigma a in - match EConstr.kind sigma (EConstr.of_constr a') with + match EConstr.kind sigma a' with | Construct cstr -> let l = try Constrmap.find (fst cstr) cstrs with Not_found -> [] in Constrmap.add (fst cstr) ((args,id)::l) cstrs @@ -691,7 +690,7 @@ let find_projectable_constructor env evd cstr k args cstr_subst = List.filter (fun (args',id) -> (* is_conv is maybe too strong (and source of useless computation) *) (* (at least expansion of aliases is needed) *) - Array.for_all2 (fun c1 c2 -> is_conv env evd c1 (EConstr.of_constr c2)) args args') l in + Array.for_all2 (fun c1 c2 -> is_conv env evd c1 c2) args args') l in List.map snd l with Not_found -> [] @@ -1104,14 +1103,14 @@ exception CannotProject of evar_map * EConstr.existential let rec is_constrainable_in top evd k (ev,(fv_rels,fv_ids) as g) t = let f,args = decompose_app_vect evd t in - match EConstr.kind evd (EConstr.of_constr f) with + match EConstr.kind evd f with | Construct ((ind,_),u) -> let n = Inductiveops.inductive_nparams ind in if n > Array.length args then true (* We don't try to be more clever *) else let params = fst (Array.chop n args) in - Array.for_all (EConstr.of_constr %> is_constrainable_in false evd k g) params - | Ind _ -> Array.for_all (EConstr.of_constr %> is_constrainable_in false evd k g) args + Array.for_all (is_constrainable_in false evd k g) params + | Ind _ -> Array.for_all (is_constrainable_in false evd k g) args | Prod (na,t1,t2) -> is_constrainable_in false evd k g t1 && is_constrainable_in false evd k g t2 | Evar (ev',_) -> top || not (Evar.equal ev' ev) (*If ev' needed, one may also try to restrict it*) | Var id -> Id.Set.mem id fv_ids @@ -1463,8 +1462,7 @@ let rec invert_definition conv_algo choose env evd pbty (evk,argsv as ev) rhs = progress := true; match let c,args = decompose_app_vect !evdref t in - let args = Array.map EConstr.of_constr args in - match EConstr.kind !evdref (EConstr.of_constr c) with + match EConstr.kind !evdref c with | Construct (cstr,u) when noccur_between !evdref 1 k t -> (* This is common case when inferring the return clause of match *) (* (currently rudimentary: we do not treat the case of multiple *) diff --git a/pretyping/inductiveops.ml b/pretyping/inductiveops.ml index cb8b25323..9c5a2e894 100644 --- a/pretyping/inductiveops.ml +++ b/pretyping/inductiveops.ml @@ -451,13 +451,13 @@ let extract_mrectype sigma t = let open EConstr in let (t, l) = decompose_app sigma t in match EConstr.kind sigma t with - | Ind ind -> (ind, List.map EConstr.Unsafe.to_constr l) + | Ind ind -> (ind, l) | _ -> raise Not_found let find_mrectype_vect env sigma c = let open EConstr in let (t, l) = Termops.decompose_app_vect sigma (EConstr.of_constr (whd_all env sigma c)) in - match EConstr.kind sigma (EConstr.of_constr t) with + match EConstr.kind sigma t with | Ind ind -> (ind, l) | _ -> raise Not_found diff --git a/pretyping/inductiveops.mli b/pretyping/inductiveops.mli index 1614e1817..4bb484759 100644 --- a/pretyping/inductiveops.mli +++ b/pretyping/inductiveops.mli @@ -161,9 +161,9 @@ val make_arity : env -> bool -> inductive_family -> sorts -> types val build_branch_type : env -> bool -> constr -> constructor_summary -> types (** Raise [Not_found] if not given a valid inductive type *) -val extract_mrectype : evar_map -> EConstr.t -> pinductive * constr list -val find_mrectype : env -> evar_map -> EConstr.types -> pinductive * constr list -val find_mrectype_vect : env -> evar_map -> EConstr.types -> pinductive * constr array +val extract_mrectype : evar_map -> EConstr.t -> pinductive * EConstr.constr list +val find_mrectype : env -> evar_map -> EConstr.types -> pinductive * EConstr.constr list +val find_mrectype_vect : env -> evar_map -> EConstr.types -> pinductive * EConstr.constr array val find_rectype : env -> evar_map -> EConstr.types -> inductive_type val find_inductive : env -> evar_map -> EConstr.types -> pinductive * constr list val find_coinductive : env -> evar_map -> EConstr.types -> pinductive * constr list diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml index c792bf2ca..f814028f9 100644 --- a/pretyping/pretyping.ml +++ b/pretyping/pretyping.ml @@ -971,7 +971,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre let pj = pretype_type empty_valcon env_p evdref lvar p in let ccl = nf_evar !evdref pj.utj_val in let pred = it_mkLambda_or_LetIn ccl psign in - let typ = lift (- nar) (EConstr.of_constr (beta_applist !evdref (pred,[cj.uj_val]))) in + let typ = lift (- nar) (beta_applist !evdref (pred,[cj.uj_val])) in pred, typ | None -> let p = match tycon with @@ -987,7 +987,6 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre let n = Context.Rel.length cs.cs_args in let pi = lift n pred in (* liftn n 2 pred ? *) let pi = beta_applist !evdref (pi, [EConstr.of_constr (build_dependent_constructor cs)]) in - let pi = EConstr.of_constr pi in let csgn = if not !allow_anonymous_refs then List.map (set_name Anonymous) cs.cs_args diff --git a/pretyping/reductionops.ml b/pretyping/reductionops.ml index 31354217f..6ec3cd985 100644 --- a/pretyping/reductionops.ml +++ b/pretyping/reductionops.ml @@ -669,7 +669,7 @@ let beta_app sigma (c,l) = let beta_applist sigma (c,l) = let zip s = Stack.zip sigma s in - EConstr.Unsafe.to_constr (stacklam zip [] sigma c (Stack.append_app_list l Stack.empty)) + stacklam zip [] sigma c (Stack.append_app_list l Stack.empty) (* Iota reduction tools *) @@ -1611,8 +1611,8 @@ let meta_reducible_instance evd b = let u = whd_betaiota Evd.empty u (** FIXME *) in let u = EConstr.of_constr u in match EConstr.kind evd u with - | Case (ci,p,c,bl) when EConstr.isMeta evd (EConstr.of_constr (strip_outer_cast evd c)) -> - let m = destMeta evd (EConstr.of_constr (strip_outer_cast evd c)) in + | Case (ci,p,c,bl) when EConstr.isMeta evd (strip_outer_cast evd c) -> + let m = destMeta evd (strip_outer_cast evd c) in (match try let g, s = Metamap.find m metas in @@ -1623,8 +1623,8 @@ let meta_reducible_instance evd b = with | Some g -> irec (mkCase (ci,p,g,bl)) | None -> mkCase (ci,irec p,c,Array.map irec bl)) - | App (f,l) when EConstr.isMeta evd (EConstr.of_constr (strip_outer_cast evd f)) -> - let m = destMeta evd (EConstr.of_constr (strip_outer_cast evd f)) in + | App (f,l) when EConstr.isMeta evd (strip_outer_cast evd f) -> + let m = destMeta evd (strip_outer_cast evd f) in (match try let g, s = Metamap.find m metas in @@ -1671,8 +1671,8 @@ let head_unfold_under_prod ts env sigma c = | Prod (n,t,c) -> mkProd (n,aux t, aux c) | _ -> let (h,l) = decompose_app_vect sigma c in - match EConstr.kind sigma (EConstr.of_constr h) with - | Const cst -> beta_app sigma (unfold cst, Array.map EConstr.of_constr l) + match EConstr.kind sigma h with + | Const cst -> beta_app sigma (unfold cst, l) | _ -> c in EConstr.Unsafe.to_constr (aux c) diff --git a/pretyping/reductionops.mli b/pretyping/reductionops.mli index 1e6527b29..3b3242537 100644 --- a/pretyping/reductionops.mli +++ b/pretyping/reductionops.mli @@ -204,7 +204,7 @@ val shrink_eta : EConstr.constr -> EConstr.constr val safe_evar_value : evar_map -> existential -> constr option -val beta_applist : evar_map -> EConstr.t * EConstr.t list -> constr +val beta_applist : evar_map -> EConstr.t * EConstr.t list -> EConstr.constr val hnf_prod_app : env -> evar_map -> EConstr.t -> EConstr.t -> constr val hnf_prod_appvect : env -> evar_map -> EConstr.t -> EConstr.t array -> constr diff --git a/pretyping/retyping.ml b/pretyping/retyping.ml index 88899e633..3142ea5cb 100644 --- a/pretyping/retyping.ml +++ b/pretyping/retyping.ml @@ -59,7 +59,7 @@ let local_def (na, b, t) = LocalDef (na, inj b, inj t) let get_type_from_constraints env sigma t = - if isEvar sigma (EConstr.of_constr (fst (decompose_app_vect sigma t))) then + if isEvar sigma (fst (decompose_app_vect sigma t)) then match List.map_filter (fun (pbty,env,t1,t2) -> if is_fconv Reduction.CONV env sigma t (EConstr.of_constr t1) then Some t2 @@ -102,7 +102,7 @@ let retype ?(polyprop=true) sigma = let rec type_of env cstr = match EConstr.kind sigma cstr with | Meta n -> - EConstr.of_constr (try strip_outer_cast sigma (EConstr.of_constr (Evd.meta_ftype sigma n).Evd.rebus) + (try strip_outer_cast sigma (EConstr.of_constr (Evd.meta_ftype sigma n).Evd.rebus) with Not_found -> retype_error (BadMeta n)) | Rel n -> let ty = EConstr.of_constr (RelDecl.get_type (lookup_rel n env)) in @@ -135,10 +135,10 @@ let retype ?(polyprop=true) sigma = | CoFix (i,(_,tys,_)) -> tys.(i) | App(f,args) when is_template_polymorphic env sigma f -> let t = type_of_global_reference_knowing_parameters env f args in - EConstr.of_constr (strip_outer_cast sigma (subst_type env sigma t (Array.to_list args))) + strip_outer_cast sigma (subst_type env sigma t (Array.to_list args)) | App(f,args) -> - EConstr.of_constr (strip_outer_cast sigma - (subst_type env sigma (type_of env f) (Array.to_list args))) + strip_outer_cast sigma + (subst_type env sigma (type_of env f) (Array.to_list args)) | Proj (p,c) -> let ty = type_of env c in EConstr.of_constr (try @@ -259,6 +259,5 @@ let expand_projection env sigma pr c args = try Inductiveops.find_mrectype env sigma (EConstr.of_constr ty) with Not_found -> retype_error BadRecursiveType in - let ind_args = List.map EConstr.of_constr ind_args in mkApp (mkConstU (Projection.constant pr,u), Array.of_list (ind_args @ (c :: args))) diff --git a/pretyping/tacred.ml b/pretyping/tacred.ml index 1ec8deb1b..1d179c683 100644 --- a/pretyping/tacred.ml +++ b/pretyping/tacred.ml @@ -427,8 +427,6 @@ let solve_arity_problem env sigma fxminargs c = let rec check strict c = let c' = EConstr.of_constr (whd_betaiotazeta sigma c) in let (h,rcargs) = decompose_app_vect sigma c' in - let rcargs = Array.map EConstr.of_constr rcargs in - let h = EConstr.of_constr h in match EConstr.kind sigma h with Evar(i,_) when Evar.Map.mem i fxminargs && not (Evd.is_defined !evm i) -> let minargs = Evar.Map.find i fxminargs in @@ -734,14 +732,13 @@ and reduce_params env sigma stack l = and whd_simpl_stack env sigma = let rec redrec s = let (x, stack) = decompose_app_vect sigma s in - let stack = Array.map_to_list EConstr.of_constr stack in - let x = EConstr.of_constr x in + let stack = Array.to_list stack in let s' = (x, stack) in match EConstr.kind sigma x with | Lambda (na,t,c) -> (match stack with | [] -> s' - | a :: rest -> redrec (EConstr.of_constr (beta_applist sigma (x, stack)))) + | a :: rest -> redrec (beta_applist sigma (x, stack))) | LetIn (n,b,t,c) -> redrec (applist (Vars.substl [b] c, stack)) | App (f,cl) -> redrec (applist(f, (Array.to_list cl)@stack)) | Cast (c,_,_) -> redrec (applist(c, stack)) @@ -1122,14 +1119,12 @@ let fold_one_com com env sigma c = unfold produces it, so that the "unfold f; fold f" configuration works to refold fix expressions *) let a = subst_term sigma (EConstr.of_constr (clos_norm_flags unfold_side_red env sigma rcom)) c in - let a = EConstr.of_constr a in if not (EConstr.eq_constr sigma a c) then Vars.subst1 com a else (* Then reason on the non beta-iota-zeta form for compatibility - even if it is probably a useless configuration *) let a = subst_term sigma rcom c in - let a = EConstr.of_constr a in Vars.subst1 com a let fold_commands cl env sigma c = @@ -1195,7 +1190,7 @@ let reduce_to_ind_gen allow_product env sigma t = let rec elimrec env t l = let t = hnf_constr env sigma t in let t = EConstr.of_constr t in - match EConstr.kind sigma (EConstr.of_constr (fst (decompose_app_vect sigma t))) with + match EConstr.kind sigma (fst (decompose_app_vect sigma t)) with | Ind ind-> (check_privacy env ind, it_mkProd_or_LetIn t l) | Prod (n,ty,t') -> let open Context.Rel.Declaration in @@ -1208,7 +1203,7 @@ let reduce_to_ind_gen allow_product env sigma t = was partially the case between V5.10 and V8.1 *) let t' = whd_all env sigma t in let t' = EConstr.of_constr t' in - match EConstr.kind sigma (EConstr.of_constr (fst (decompose_app_vect sigma t'))) with + match EConstr.kind sigma (fst (decompose_app_vect sigma t')) with | Ind ind-> (check_privacy env ind, it_mkProd_or_LetIn t' l) | _ -> user_err (str"Not an inductive product.") in @@ -1275,7 +1270,6 @@ let reduce_to_ref_gen allow_product env sigma ref t = (* lazily reduces to match the head of [t] with the expected [ref] *) let rec elimrec env t l = let c, _ = decompose_app_vect sigma t in - let c = EConstr.of_constr c in match EConstr.kind sigma c with | Prod (n,ty,t') -> if allow_product then diff --git a/pretyping/typeclasses.ml b/pretyping/typeclasses.ml index a970c434f..f59a6dcd9 100644 --- a/pretyping/typeclasses.ml +++ b/pretyping/typeclasses.ml @@ -156,17 +156,17 @@ let class_of_constr sigma c = try Some (dest_class_arity (Global.env ()) sigma c) with e when CErrors.noncritical e -> None -let is_class_constr c = - try let gr, u = Universes.global_of_constr c in +let is_class_constr sigma c = + try let gr, u = Termops.global_of_constr sigma c in Refmap.mem gr !classes with Not_found -> false let rec is_class_type evd c = let c, _ = Termops.decompose_app_vect evd c in - match EConstr.kind evd (EConstr.of_constr c) with + match EConstr.kind evd c with | Prod (_, _, t) -> is_class_type evd t | Cast (t, _, _) -> is_class_type evd t - | _ -> is_class_constr c + | _ -> is_class_constr evd c let is_class_evar evd evi = is_class_type evd (EConstr.of_constr evi.Evd.evar_concl) diff --git a/pretyping/typing.ml b/pretyping/typing.ml index 29697260f..40ef2450a 100644 --- a/pretyping/typing.ml +++ b/pretyping/typing.ml @@ -138,7 +138,7 @@ let e_type_case_branches env evdref (ind,largs) pj c = let nparams = inductive_params specif in let (params,realargs) = List.chop nparams largs in let p = pj.uj_val in - let realargs = List.map EConstr.of_constr realargs in + let params = List.map EConstr.Unsafe.to_constr params in let () = e_is_correct_arity env evdref c pj ind specif params in let lc = build_branches_type ind specif params (EConstr.to_constr !evdref p) in let lc = Array.map EConstr.of_constr lc in @@ -232,7 +232,6 @@ let judge_of_projection env sigma p cj = try find_mrectype env sigma cj.uj_type with Not_found -> error_case_not_inductive env sigma cj in - let args = List.map EConstr.of_constr args in let ty = EConstr.of_constr (CVars.subst_instance_constr u pb.Declarations.proj_type) in let ty = substl (cj.uj_val :: List.rev args) ty in {uj_val = EConstr.mkProj (p,cj.uj_val); diff --git a/pretyping/unification.ml b/pretyping/unification.ml index 81d9ecad5..169dd45bc 100644 --- a/pretyping/unification.ml +++ b/pretyping/unification.ml @@ -632,7 +632,7 @@ let check_compatibility env pbty flags (sigma,metasubst,evarsubst : subst0) tyM let rec is_neutral env sigma ts t = let (f, l) = decompose_app_vect sigma t in - match EConstr.kind sigma (EConstr.of_constr f) with + match EConstr.kind sigma f with | Const (c, u) -> not (Environ.evaluable_constant c env) || not (is_transparent env (ConstKey c)) || @@ -1488,10 +1488,6 @@ let w_unify_core_0 env evd with_types cv_pb flags m n = let w_typed_unify env evd = w_unify_core_0 env evd true let w_typed_unify_array env evd flags f1 l1 f2 l2 = - let f1 = EConstr.of_constr f1 in - let f2 = EConstr.of_constr f2 in - let l1 = Array.map EConstr.of_constr l1 in - let l2 = Array.map EConstr.of_constr l2 in let f1,l1,f2,l2 = adjust_app_array_size f1 l1 f2 l2 in let (mc1,evd') = retract_coercible_metas evd in let fold_subst subst m n = unify_0_with_initial_metas subst true env CONV flags.core_unify_flags m n in @@ -1743,7 +1739,7 @@ let w_unify_to_subterm env evd ?(flags=default_unify_flags ()) (op,cl) = let bestexn = ref None in let kop = Keys.constr_key (EConstr.to_constr evd op) in let rec matchrec cl = - let cl = EConstr.of_constr (strip_outer_cast evd cl) in + let cl = strip_outer_cast evd cl in (try if closed0 evd cl && not (isEvar evd cl) && keyed_unify env evd kop cl then (try @@ -1837,7 +1833,6 @@ let w_unify_to_subterm_all env evd ?(flags=default_unify_flags ()) (op,cl) = in let rec matchrec cl = let cl = strip_outer_cast evd cl in - let cl = EConstr.of_constr cl in (bind (if closed0 evd cl then return (fun () -> w_typed_unify env evd CONV flags op cl,cl) @@ -1898,7 +1893,7 @@ let w_unify_to_subterm_list env evd flags hdmeta oplist t = unify pre-existing non frozen evars of the goal or of the pattern *) set_no_delta_flags flags in - let t' = (EConstr.of_constr (strip_outer_cast evd op),t) in + let t' = (strip_outer_cast evd op,t) in let (evd',cl) = try if is_keyed_unification () then @@ -1992,7 +1987,7 @@ let w_unify env evd cv_pb ?(flags=default_unify_flags ()) ty1 ty2 = let hd2,l2 = decompose_app_vect evd (EConstr.of_constr (whd_nored evd ty2)) in let is_empty1 = Array.is_empty l1 in let is_empty2 = Array.is_empty l2 in - match kind_of_term hd1, not is_empty1, kind_of_term hd2, not is_empty2 with + match EConstr.kind evd hd1, not is_empty1, EConstr.kind evd hd2, not is_empty2 with (* Pattern case *) | (Meta _, true, Lambda _, _ | Lambda _, _, Meta _, true) when Int.equal (Array.length l1) (Array.length l2) -> diff --git a/proofs/clenv.ml b/proofs/clenv.ml index d98669660..572db1d40 100644 --- a/proofs/clenv.ml +++ b/proofs/clenv.ml @@ -266,7 +266,7 @@ let clenv_unify_meta_types ?(flags=default_unify_flags ()) clenv = let clenv_unique_resolver ?(flags=default_unify_flags ()) clenv gl = let concl = Goal.V82.concl clenv.evd (sig_it gl) in - if isMeta clenv.evd (EConstr.of_constr (fst (decompose_app_vect clenv.evd (EConstr.of_constr (whd_nored clenv.evd clenv.templtyp.rebus))))) then + if isMeta clenv.evd (fst (decompose_app_vect clenv.evd (EConstr.of_constr (whd_nored clenv.evd clenv.templtyp.rebus)))) then clenv_unify CUMUL ~flags (clenv_type clenv) concl (clenv_unify_meta_types ~flags clenv) else @@ -480,7 +480,7 @@ let error_already_defined b = (str "Position " ++ int n ++ str" already defined.") let clenv_unify_binding_type clenv c t u = - if isMeta clenv.evd (EConstr.of_constr (fst (decompose_app_vect clenv.evd (EConstr.of_constr (whd_nored clenv.evd u))))) then + if isMeta clenv.evd (fst (decompose_app_vect clenv.evd (EConstr.of_constr (whd_nored clenv.evd u)))) then (* Not enough information to know if some subtyping is needed *) CoerceToType, clenv, c else diff --git a/proofs/clenvtac.ml b/proofs/clenvtac.ml index 539b1bcb2..84cdc09ee 100644 --- a/proofs/clenvtac.ml +++ b/proofs/clenvtac.ml @@ -34,7 +34,7 @@ let clenv_cast_meta clenv = | _ -> EConstr.map clenv.evd crec u and crec_hd u = - match EConstr.kind clenv.evd (EConstr.of_constr (strip_outer_cast clenv.evd u)) with + match EConstr.kind clenv.evd (strip_outer_cast clenv.evd u) with | Meta mv -> (try let b = Typing.meta_type clenv.evd mv in diff --git a/tactics/class_tactics.ml b/tactics/class_tactics.ml index 0ca661557..6c45bef1c 100644 --- a/tactics/class_tactics.ml +++ b/tactics/class_tactics.ml @@ -1496,7 +1496,6 @@ let _ = let rec head_of_constr sigma t = let t = strip_outer_cast sigma (EConstr.of_constr (collapse_appl sigma t)) in - let t = EConstr.of_constr t in match EConstr.kind sigma t with | Prod (_,_,c2) -> head_of_constr sigma c2 | LetIn (_,_,_,c2) -> head_of_constr sigma c2 diff --git a/tactics/equality.ml b/tactics/equality.ml index 4c79a6199..2eead2d22 100644 --- a/tactics/equality.ml +++ b/tactics/equality.ml @@ -183,7 +183,7 @@ let instantiate_lemma_all frzevars gl c ty l l2r concl = let flags = make_flags frzevars (Tacmach.New.project gl) rewrite_unif_flags eqclause in let occs = w_unify_to_subterm_all ~flags env eqclause.evd - (EConstr.of_constr (if l2r then c1 else c2),EConstr.of_constr concl) + ((if l2r then c1 else c2),concl) in List.map try_occ occs let instantiate_lemma gl c ty l l2r concl = @@ -314,6 +314,7 @@ let general_elim_clause with_evars frzevars tac cls c t l l2r elim = | None -> pf_nf_concl gl | Some id -> pf_get_hyp_typ id (Proofview.Goal.assume gl) in + let typ = EConstr.of_constr typ in let cs = instantiate_lemma typ in if firstonly then tclFIRST (List.map try_clause cs) else tclMAP try_clause cs @@ -1207,7 +1208,6 @@ let sig_clausal_form env sigma sort_of_ty siglen ty dflt = | _ -> anomaly ~label:"sig_clausal_form" (Pp.str "should be a sigma type") in let ev = Evarutil.e_new_evar env evdref a in let rty = beta_applist sigma (p_i_minus_1,[ev]) in - let rty = EConstr.of_constr rty in let tuple_tail = sigrec_clausal_form (siglen-1) rty in let evopt = match EConstr.kind !evdref ev with Evar _ -> None | _ -> Some ev in match evopt with @@ -1348,10 +1348,6 @@ let inject_if_homogenous_dependent_pair ty = if not (Termops.is_global sigma (sigTconstr()) eqTypeDest) then raise Exit; let hd1,ar1 = decompose_app_vect sigma t1 and hd2,ar2 = decompose_app_vect sigma t2 in - let hd1 = EConstr.of_constr hd1 in - let hd2 = EConstr.of_constr hd2 in - let ar1 = Array.map EConstr.of_constr ar1 in - let ar2 = Array.map EConstr.of_constr ar2 in if not (Termops.is_global sigma (existTconstr()) hd1) then raise Exit; if not (Termops.is_global sigma (existTconstr()) hd2) then raise Exit; let ind,_ = try pf_apply find_mrectype gl ar1.(0) with Not_found -> raise Exit in @@ -1565,7 +1561,6 @@ let decomp_tuple_term env sigma c t = let car_code = applist (mkConstU (destConstRef p1,i),[a;p;inner_code]) and cdr_code = applist (mkConstU (destConstRef p2,i),[a;p;inner_code]) in let cdrtyp = beta_applist sigma (p,[car]) in - let cdrtyp = EConstr.of_constr cdrtyp in List.map (fun l -> ((car,a),car_code)::l) (decomprec cdr_code cdr cdrtyp) with Constr_matching.PatternMatchingFailure -> [] @@ -1593,13 +1588,11 @@ let subst_tuple_term env sigma dep_pair1 dep_pair2 b = (* We build the expected goal *) let abst_B = List.fold_right - (fun (e,t) body -> lambda_create env (t,EConstr.of_constr (subst_term sigma e body))) e1_list b in + (fun (e,t) body -> lambda_create env (t,subst_term sigma e body)) e1_list b in let pred_body = beta_applist sigma (abst_B,proj_list) in - let pred_body = EConstr.of_constr pred_body in let body = mkApp (lambda_create env (typ,pred_body),[|dep_pair1|]) in let expected_goal = beta_applist sigma (abst_B,List.map fst e2_list) in (* Simulate now the normalisation treatment made by Logic.mk_refgoals *) - let expected_goal = EConstr.of_constr expected_goal in let expected_goal = nf_betaiota sigma expected_goal in let expected_goal = EConstr.of_constr expected_goal in (* Retype to get universes right *) diff --git a/tactics/hints.ml b/tactics/hints.ml index cd5fc79f5..2b310033c 100644 --- a/tactics/hints.ml +++ b/tactics/hints.ml @@ -50,7 +50,6 @@ exception Bound let head_constr_bound sigma t = let t = strip_outer_cast sigma t in - let t = EConstr.of_constr t in let _,ccl = decompose_prod_assum sigma t in let hd,args = decompose_app sigma ccl in match EConstr.kind sigma hd with @@ -66,11 +65,8 @@ let head_constr sigma c = let decompose_app_bound sigma t = let t = strip_outer_cast sigma t in - let t = EConstr.of_constr t in let _,ccl = decompose_prod_assum sigma t in let hd,args = decompose_app_vect sigma ccl in - let hd = EConstr.of_constr hd in - let args = Array.map EConstr.of_constr args in match EConstr.kind sigma hd with | Const (c,u) -> ConstRef c, args | Ind (i,u) -> IndRef i, args @@ -754,7 +750,6 @@ let secvars_of_global env gr = let make_exact_entry env sigma pri poly ?(name=PathAny) (c, cty, ctx) = let secvars = secvars_of_constr env sigma c in let cty = strip_outer_cast sigma cty in - let cty = EConstr.of_constr cty in match EConstr.kind sigma cty with | Prod _ -> failwith "make_exact_entry" | _ -> diff --git a/tactics/inv.ml b/tactics/inv.ml index 5c296b343..ac9a564e5 100644 --- a/tactics/inv.ml +++ b/tactics/inv.ml @@ -462,7 +462,6 @@ let raw_inversion inv_kind id status names = Reductionops.beta_applist sigma (elim_predicate, realargs), case_nodep_then_using in - let cut_concl = EConstr.of_constr cut_concl in let refined id = let prf = mkApp (mkVar id, args) in Refine.refine { run = fun h -> Sigma (prf, h, Sigma.refl) } diff --git a/tactics/tactics.ml b/tactics/tactics.ml index 606eb27b9..03f81773b 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -528,7 +528,7 @@ fun env sigma p -> function let Sigma (rem, sigma, r) = mk_holes env sigma (p +> q) rem in Sigma (arg :: rem, sigma, r) -let rec check_mutind env sigma k cl = match EConstr.kind sigma (EConstr.of_constr (strip_outer_cast sigma cl)) with +let rec check_mutind env sigma k cl = match EConstr.kind sigma (strip_outer_cast sigma cl) with | Prod (na, c1, b) -> if Int.equal k 1 then try @@ -1647,10 +1647,8 @@ let make_projection env sigma params cstr sign elim i n c u = then let t = lift (i+1-n) t in let abselim = beta_applist sigma (elim, params@[t;branch]) in - let abselim = EConstr.of_constr abselim in let args = Array.map EConstr.of_constr (Context.Rel.to_extended_vect 0 sign) in let c = beta_applist sigma (abselim, [mkApp (c, args)]) in - let c = EConstr.of_constr c in Some (it_mkLambda_or_LetIn c sign, it_mkProd_or_LetIn t sign) else None @@ -2856,7 +2854,7 @@ let generalize_goal_gen env sigma ids i ((occs,c,b),na) t cl = let open Context.Rel.Declaration in let decls,cl = decompose_prod_n_assum sigma i cl in let dummy_prod = it_mkProd_or_LetIn mkProp decls in - let newdecls,_ = decompose_prod_n_assum sigma i (EConstr.of_constr (subst_term_gen sigma EConstr.eq_constr_nounivs c dummy_prod)) in + let newdecls,_ = decompose_prod_n_assum sigma i (subst_term_gen sigma EConstr.eq_constr_nounivs c dummy_prod) in let cl',sigma' = subst_closed_term_occ env sigma (AtOccs occs) c (it_mkProd_or_LetIn cl newdecls) in let cl' = EConstr.of_constr cl' in let na = generalized_name sigma c t ids cl' na in diff --git a/toplevel/class.ml b/toplevel/class.ml index 6788cf1b7..e55489471 100644 --- a/toplevel/class.ml +++ b/toplevel/class.ml @@ -81,12 +81,13 @@ let check_target clt = function (* condition d'heritage uniforme *) -let uniform_cond nargs lt = +let uniform_cond sigma nargs lt = + let open EConstr in let rec aux = function | (0,[]) -> true | (n,t::l) -> - let t = strip_outer_cast Evd.empty t (** FIXME *) in - isRel t && Int.equal (destRel t) n && aux ((n-1),l) + let t = strip_outer_cast sigma t in + isRel sigma t && Int.equal (destRel sigma t) n && aux ((n-1),l) | _ -> false in aux (nargs,lt) @@ -263,7 +264,7 @@ let add_new_coercion_core coef stre poly source target isid = raise (CoercionError (NoSource source)) in check_source (Some cls); - if not (uniform_cond (llp-ind) lvs) then + if not (uniform_cond Evd.empty (** FIXME *) (llp-ind) lvs) then warn_uniform_inheritance coef; let clt = try diff --git a/toplevel/obligations.ml b/toplevel/obligations.ml index 8aee9d241..aac1d1ed4 100644 --- a/toplevel/obligations.ml +++ b/toplevel/obligations.ml @@ -396,7 +396,7 @@ let subst_deps expand obls deps t = (Vars.replace_vars (List.map (fun (n, (_, b)) -> n, b) osubst) t) let rec prod_app t n = - match kind_of_term (Termops.strip_outer_cast Evd.empty (EConstr.of_constr t)) (** FIXME *) with + match kind_of_term (EConstr.Unsafe.to_constr (Termops.strip_outer_cast Evd.empty (EConstr.of_constr t))) (** FIXME *) with | Prod (_,_,b) -> subst1 n b | LetIn (_, b, t, b') -> prod_app (subst1 b b') n | _ -> diff --git a/toplevel/search.ml b/toplevel/search.ml index 653d4ac5c..102bd003d 100644 --- a/toplevel/search.ml +++ b/toplevel/search.ml @@ -111,20 +111,20 @@ let generic_search glnumopt fn = (** This function tries to see whether the conclusion matches a pattern. *) (** FIXME: this is quite dummy, we may find a more efficient algorithm. *) -let rec pattern_filter pat ref env typ = - let typ = Termops.strip_outer_cast Evd.empty (EConstr.of_constr typ) (** FIXME *) in - if Constr_matching.is_matching env Evd.empty pat (EConstr.of_constr typ) then true - else match kind_of_term typ with +let rec pattern_filter pat ref env sigma typ = + let typ = Termops.strip_outer_cast sigma typ in + if Constr_matching.is_matching env sigma pat typ then true + else match EConstr.kind sigma typ with | Prod (_, _, typ) - | LetIn (_, _, _, typ) -> pattern_filter pat ref env typ + | LetIn (_, _, _, typ) -> pattern_filter pat ref env sigma typ | _ -> false -let rec head_filter pat ref env typ = - let typ = Termops.strip_outer_cast Evd.empty (EConstr.of_constr typ) (** FIXME *) in - if Constr_matching.is_matching_head env Evd.empty pat (EConstr.of_constr typ) then true - else match kind_of_term typ with +let rec head_filter pat ref env sigma typ = + let typ = Termops.strip_outer_cast sigma typ in + if Constr_matching.is_matching_head env sigma pat typ then true + else match EConstr.kind sigma typ with | Prod (_, _, typ) - | LetIn (_, _, _, typ) -> head_filter pat ref env typ + | LetIn (_, _, _, typ) -> head_filter pat ref env sigma typ | _ -> false let full_name_of_reference ref = @@ -162,7 +162,7 @@ let search_pattern gopt pat mods pr_search = let blacklist_filter = blacklist_filter_aux () in let filter ref env typ = module_filter mods ref env typ && - pattern_filter pat ref env typ && + pattern_filter pat ref env Evd.empty (* FIXME *) (EConstr.of_constr typ) && blacklist_filter ref env typ in let iter ref env typ = @@ -186,8 +186,8 @@ let search_rewrite gopt pat mods pr_search = let blacklist_filter = blacklist_filter_aux () in let filter ref env typ = module_filter mods ref env typ && - (pattern_filter pat1 ref env typ || - pattern_filter pat2 ref env typ) && + (pattern_filter pat1 ref env Evd.empty (* FIXME *) (EConstr.of_constr typ) || + pattern_filter pat2 ref env Evd.empty (EConstr.of_constr typ)) && blacklist_filter ref env typ in let iter ref env typ = @@ -201,7 +201,7 @@ let search_by_head gopt pat mods pr_search = let blacklist_filter = blacklist_filter_aux () in let filter ref env typ = module_filter mods ref env typ && - head_filter pat ref env typ && + head_filter pat ref env Evd.empty (* FIXME *) (EConstr.of_constr typ) && blacklist_filter ref env typ in let iter ref env typ = diff --git a/toplevel/vernacentries.ml b/toplevel/vernacentries.ml index 3b8c62738..773766d7e 100644 --- a/toplevel/vernacentries.ml +++ b/toplevel/vernacentries.ml @@ -103,11 +103,12 @@ let try_print_subgoals () = (* Simulate the Intro(s) tactic *) let show_intro all = + let open EConstr in let pf = get_pftreestate() in let {Evd.it=gls ; sigma=sigma; } = Proof.V82.subgoals pf in if not (List.is_empty gls) then begin let gl = {Evd.it=List.hd gls ; sigma = sigma; } in - let l,_= decompose_prod_assum (Termops.strip_outer_cast (project gl) (EConstr.of_constr (pf_concl gl))) in + let l,_= decompose_prod_assum sigma (Termops.strip_outer_cast sigma (EConstr.of_constr (pf_concl gl))) in if all then let lid = Tactics.find_intro_names l gl in Feedback.msg_notice (hov 0 (prlist_with_sep spc pr_id lid)) -- cgit v1.2.3 From e09f3b44bb381854b647a6d9debdeddbfc49177e Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sun, 20 Nov 2016 22:16:08 +0100 Subject: Proofview.Goal primitive now return EConstrs. --- engine/evarutil.ml | 1 + engine/evarutil.mli | 2 +- engine/proofview.ml | 4 +- engine/proofview.mli | 4 +- ltac/extratactics.ml4 | 11 +++-- ltac/g_class.ml4 | 2 - ltac/rewrite.ml | 8 +--- ltac/tacinterp.ml | 1 - ltac/tactic_debug.ml | 1 + plugins/btauto/refl_btauto.ml | 2 + plugins/cc/cctac.ml | 5 ++- plugins/decl_mode/decl_proof_instr.ml | 2 +- plugins/fourier/fourierR.ml | 2 +- plugins/omega/coq_omega.ml | 1 + plugins/quote/quote.ml | 1 + plugins/ssrmatching/ssrmatching.ml4 | 2 +- proofs/clenvtac.ml | 2 +- proofs/refine.ml | 2 - proofs/tacmach.ml | 8 ++-- proofs/tacmach.mli | 10 ++--- tactics/auto.ml | 2 - tactics/class_tactics.ml | 4 +- tactics/eauto.ml | 4 +- tactics/eqdecide.ml | 2 - tactics/equality.ml | 10 ++--- tactics/hipattern.ml | 1 - tactics/inv.ml | 7 ++-- tactics/leminv.ml | 2 +- tactics/tacticals.ml | 2 +- tactics/tactics.ml | 75 ++++++++++++----------------------- toplevel/auto_ind_decl.ml | 21 +++++----- 31 files changed, 83 insertions(+), 118 deletions(-) diff --git a/engine/evarutil.ml b/engine/evarutil.ml index 204997445..73286f2c4 100644 --- a/engine/evarutil.ml +++ b/engine/evarutil.ml @@ -611,6 +611,7 @@ let clear_hyps_in_evi_main env evdref hyps terms ids = (nhyps,terms) let clear_hyps_in_evi env evdref hyps concl ids = + let concl = EConstr.Unsafe.to_constr concl in match clear_hyps_in_evi_main env evdref hyps [concl] ids with | (nhyps,[nconcl]) -> (nhyps,nconcl) | _ -> assert false diff --git a/engine/evarutil.mli b/engine/evarutil.mli index cf36f5953..5882f02b9 100644 --- a/engine/evarutil.mli +++ b/engine/evarutil.mli @@ -199,7 +199,7 @@ exception ClearDependencyError of Id.t * clear_dependency_error used by [Goal] and (indirectly) [Proofview] to handle the clear tactic gracefully*) val cleared : bool Store.field -val clear_hyps_in_evi : env -> evar_map ref -> named_context_val -> types -> +val clear_hyps_in_evi : env -> evar_map ref -> named_context_val -> EConstr.types -> Id.Set.t -> named_context_val * types val clear_hyps2_in_evi : env -> evar_map ref -> named_context_val -> types -> types -> diff --git a/engine/proofview.ml b/engine/proofview.ml index 9adf94744..9e5e9c7da 100644 --- a/engine/proofview.ml +++ b/engine/proofview.ml @@ -984,7 +984,7 @@ module Goal = struct type ('a, 'r) t = { env : Environ.env; sigma : Evd.evar_map; - concl : Term.constr ; + concl : EConstr.constr ; self : Evar.t ; (* for compatibility with old-style definitions *) } @@ -1005,7 +1005,7 @@ module Goal = struct 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 = diff --git a/engine/proofview.mli b/engine/proofview.mli index 725445251..2350592a2 100644 --- a/engine/proofview.mli +++ b/engine/proofview.mli @@ -484,7 +484,7 @@ module Goal : sig 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 concl : ([ `NF ], 'r) t -> EConstr.constr val hyps : ([ `NF ], 'r) t -> Context.Named.t val env : ('a, 'r) t -> Environ.env val sigma : ('a, 'r) t -> 'r Sigma.t @@ -492,7 +492,7 @@ module Goal : sig (** Returns the goal's conclusion even if the goal is not normalised. *) - val raw_concl : ('a, 'r) t -> Term.constr + val raw_concl : ('a, 'r) t -> EConstr.constr type ('a, 'b) enter = { enter : 'r. ('a, 'r) t -> 'b } diff --git a/ltac/extratactics.ml4 b/ltac/extratactics.ml4 index d53248a04..188d041c1 100644 --- a/ltac/extratactics.ml4 +++ b/ltac/extratactics.ml4 @@ -355,7 +355,7 @@ let refine_tac ist simple c = let concl = Proofview.Goal.concl gl in let env = Proofview.Goal.env gl in let flags = constr_flags in - let expected_type = Pretyping.OfType (EConstr.of_constr concl) in + let expected_type = Pretyping.OfType concl in let c = Pretyping.type_uconstr ~flags ~expected_type ist c in let update = { run = fun sigma -> let Sigma (c, sigma, p) = c.delayed env sigma in @@ -647,7 +647,6 @@ let hResolve id c occ t = let sigma = Sigma.to_evar_map sigma in let env = Termops.clear_named_body id (Proofview.Goal.env gl) in let concl = Proofview.Goal.concl gl in - let concl = EConstr.of_constr concl in let env_ids = Termops.ids_of_context env in let c_raw = Detyping.detype true env_ids env sigma c in let t_raw = Detyping.detype true env_ids env sigma t in @@ -694,6 +693,7 @@ let hget_evar n = Proofview.Goal.nf_enter { enter = begin fun gl -> let sigma = Tacmach.New.project gl in let concl = Proofview.Goal.concl gl in + let concl = EConstr.Unsafe.to_constr concl in let evl = evar_list concl in let concl = EConstr.of_constr concl in if List.length evl < n then @@ -748,7 +748,6 @@ let mkCaseEq a : unit Proofview.tactic = [Tactics.generalize [(mkApp(EConstr.of_constr (delayed_force refl_equal), [| type_of_a; a|]))]; Proofview.Goal.nf_enter { enter = begin fun gl -> let concl = Proofview.Goal.concl gl in - let concl = EConstr.of_constr concl in let env = Proofview.Goal.env gl in (** FIXME: this looks really wrong. Does anybody really use this tactic? *) let Sigma (c, _, _) = (Tacred.pattern_occs [Locus.OnlyOccurrences [1], a]).Reductionops.e_redfun env (Sigma.Unsafe.of_evar_map Evd.empty) concl in @@ -761,14 +760,14 @@ let mkCaseEq a : unit Proofview.tactic = let case_eq_intros_rewrite x = Proofview.Goal.nf_enter { enter = begin fun gl -> - let n = nb_prod (Tacmach.New.project gl) (EConstr.of_constr (Proofview.Goal.concl gl)) in + let n = nb_prod (Tacmach.New.project gl) (Proofview.Goal.concl gl) in (* Pp.msgnl (Printer.pr_lconstr x); *) Tacticals.New.tclTHENLIST [ mkCaseEq x; Proofview.Goal.nf_enter { enter = begin fun gl -> let concl = Proofview.Goal.concl gl in let hyps = Tacmach.New.pf_ids_of_hyps gl in - let n' = nb_prod (Tacmach.New.project gl) (EConstr.of_constr concl) in + let n' = nb_prod (Tacmach.New.project gl) concl in let h = Tacmach.New.of_old (fun g -> fresh_id hyps (Id.of_string "heq") g) gl in Tacticals.New.tclTHENLIST [ Tacticals.New.tclDO (n'-n-1) intro; @@ -809,7 +808,7 @@ let destauto_in id = end } TACTIC EXTEND destauto -| [ "destauto" ] -> [ Proofview.Goal.nf_enter { enter = begin fun gl -> destauto (EConstr.of_constr (Proofview.Goal.concl gl)) end } ] +| [ "destauto" ] -> [ Proofview.Goal.nf_enter { enter = begin fun gl -> destauto (Proofview.Goal.concl gl) end } ] | [ "destauto" "in" hyp(id) ] -> [ destauto_in id ] END diff --git a/ltac/g_class.ml4 b/ltac/g_class.ml4 index d4b84284f..a983a4fea 100644 --- a/ltac/g_class.ml4 +++ b/ltac/g_class.ml4 @@ -92,12 +92,10 @@ let rec eq_constr_mod_evars sigma x y = let progress_evars t = Proofview.Goal.nf_enter { enter = begin fun gl -> let concl = Proofview.Goal.concl gl in - let concl = EConstr.of_constr concl in let check = Proofview.Goal.nf_enter { enter = begin fun gl' -> let sigma = Tacmach.New.project gl' in let newconcl = Proofview.Goal.concl gl' in - let newconcl = EConstr.of_constr newconcl in if eq_constr_mod_evars sigma concl newconcl then Tacticals.New.tclFAIL 0 (Pp.str"No progress made (modulo evars)") else Proofview.tclUNIT () diff --git a/ltac/rewrite.ml b/ltac/rewrite.ml index 119e872f8..4c350b093 100644 --- a/ltac/rewrite.ml +++ b/ltac/rewrite.ml @@ -1595,7 +1595,6 @@ let rec insert_dependent env sigma decl accu hyps = match hyps with let assert_replacing id newt tac = let prf = Proofview.Goal.nf_enter { enter = begin fun gl -> let concl = Proofview.Goal.concl gl in - let concl = EConstr.of_constr concl in let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in let ctx = Environ.named_context env in @@ -1665,7 +1664,6 @@ let cl_rewrite_clause_newtac ?abs ?origsigma ~progress strat clause = in Proofview.Goal.nf_enter { enter = begin fun gl -> let concl = Proofview.Goal.concl gl in - let concl = EConstr.of_constr concl in let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in let ty = match clause with @@ -2143,10 +2141,9 @@ let get_hyp gl (c,l) clause l2r = let env = Tacmach.New.pf_env gl in let sigma, hi = decompose_applied_relation env evars (c,l) in let but = match clause with - | Some id -> Tacmach.New.pf_get_hyp_typ id gl - | None -> Evarutil.nf_evar evars (Tacmach.New.pf_concl gl) + | Some id -> EConstr.of_constr (Tacmach.New.pf_get_hyp_typ id gl) + | None -> nf_evar evars (Tacmach.New.pf_concl gl) in - let but = EConstr.of_constr but in unification_rewrite l2r hi.c1 hi.c2 sigma hi.prf hi.car hi.rel but env let general_rewrite_flags = { under_lambdas = false; on_morphisms = true } @@ -2195,7 +2192,6 @@ let setoid_proof ty fn fallback = let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in let concl = Proofview.Goal.concl gl in - let concl = EConstr.of_constr concl in Proofview.tclORELSE begin try diff --git a/ltac/tacinterp.ml b/ltac/tacinterp.ml index 3ed40357e..572fa32b7 100644 --- a/ltac/tacinterp.ml +++ b/ltac/tacinterp.ml @@ -1536,7 +1536,6 @@ and interp_match_goal ist lz lr lmr = let hyps = Proofview.Goal.hyps gl in let hyps = if lr then List.rev hyps else hyps in let concl = Proofview.Goal.concl gl in - let concl = EConstr.of_constr concl in let ilr = read_match_rule (extract_ltac_constr_values ist env) ist env sigma lmr in interp_match_successes lz ist (Tactic_matching.match_goal env sigma hyps concl ilr) end } diff --git a/ltac/tactic_debug.ml b/ltac/tactic_debug.ml index 5cbddc7f6..be1123d5c 100644 --- a/ltac/tactic_debug.ml +++ b/ltac/tactic_debug.ml @@ -51,6 +51,7 @@ let db_pr_goal gl = let env = Proofview.Goal.env gl in let concl = Proofview.Goal.concl gl in let penv = print_named_context env in + let concl = EConstr.Unsafe.to_constr concl in let pc = print_constr_env env concl in str" " ++ hv 0 (penv ++ fnl () ++ str "============================" ++ fnl () ++ diff --git a/plugins/btauto/refl_btauto.ml b/plugins/btauto/refl_btauto.ml index 4d2859fb0..93bd88fe4 100644 --- a/plugins/btauto/refl_btauto.ml +++ b/plugins/btauto/refl_btauto.ml @@ -221,6 +221,7 @@ module Btauto = struct Proofview.Goal.nf_enter { enter = begin fun gl -> let concl = Proofview.Goal.concl gl in let eq = Lazy.force eq in + let concl = EConstr.Unsafe.to_constr concl in let t = decomp_term (Tacmach.New.project gl) concl in match t with | Term.App (c, [|typ; p; _|]) when c === eq -> @@ -235,6 +236,7 @@ module Btauto = struct let tac = Proofview.Goal.nf_enter { enter = begin fun gl -> let concl = Proofview.Goal.concl gl in + let concl = EConstr.Unsafe.to_constr concl in let sigma = Tacmach.New.project gl in let eq = Lazy.force eq in let bool = Lazy.force Bool.typ in diff --git a/plugins/cc/cctac.ml b/plugins/cc/cctac.ml index 6295e004e..130f01e97 100644 --- a/plugins/cc/cctac.ml +++ b/plugins/cc/cctac.ml @@ -230,7 +230,7 @@ let make_prb gls depth additionnal_terms = let build_projection intype (cstr:pconstructor) special default gls= let ci= (snd(fst cstr)) in - let body=Equality.build_selector (pf_env gls) (project gls) ci (EConstr.mkRel 1) (EConstr.of_constr intype) (EConstr.of_constr special) (EConstr.of_constr default) in + let body=Equality.build_selector (pf_env gls) (project gls) ci (EConstr.mkRel 1) (EConstr.of_constr intype) (EConstr.of_constr special) default in let body = EConstr.Unsafe.to_constr body in let id=pf_get_new_id (Id.of_string "t") gls in mkLambda(Name id,intype,body) @@ -321,6 +321,7 @@ let rec proof_tac p : unit Proofview.tactic = let special=mkRel (1+nargs-argind) in refresh_universes (type_of ti) (fun intype -> refresh_universes (type_of default) (fun outtype -> + let default = EConstr.of_constr default in let proj = Tacmach.New.of_old (build_projection intype cstr special default) gl in @@ -388,6 +389,7 @@ let discriminate_tac (cstr,u as cstru) p = let pred = mkLambda(Name xid,outtype,mkRel 1) in let hid = Tacmach.New.of_old (pf_get_new_id (Id.of_string "Heq")) gl in let proj = Tacmach.New.of_old (build_projection intype cstru trivial concl) gl in + let concl = EConstr.Unsafe.to_constr concl in let injt=app_global _f_equal [|intype;outtype;proj;t1;t2;mkVar hid|] in let endt k = @@ -498,6 +500,7 @@ let mk_eq f c1 c2 k = let f_equal = Proofview.Goal.nf_enter { enter = begin fun gl -> let concl = Proofview.Goal.concl gl in + let concl = EConstr.Unsafe.to_constr concl in let cut_eq c1 c2 = try (* type_of can raise an exception *) Tacticals.New.tclTHENS diff --git a/plugins/decl_mode/decl_proof_instr.ml b/plugins/decl_mode/decl_proof_instr.ml index e73166be2..c3254010a 100644 --- a/plugins/decl_mode/decl_proof_instr.ml +++ b/plugins/decl_mode/decl_proof_instr.ml @@ -41,7 +41,7 @@ let clear ids { it = goal; sigma } = let ids = List.fold_left (fun accu x -> Id.Set.add x accu) Id.Set.empty ids in let env = Goal.V82.env sigma goal in let sign = Goal.V82.hyps sigma goal in - let cl = EConstr.Unsafe.to_constr (Goal.V82.concl sigma goal) in + let cl = Goal.V82.concl sigma goal in let evdref = ref (Evd.clear_metas sigma) in let (hyps, concl) = try Evarutil.clear_hyps_in_evi env evdref sign cl ids diff --git a/plugins/fourier/fourierR.ml b/plugins/fourier/fourierR.ml index f9802ee5e..bffca6223 100644 --- a/plugins/fourier/fourierR.ml +++ b/plugins/fourier/fourierR.ml @@ -470,7 +470,7 @@ let rec fourier () = let concl = Proofview.Goal.concl gl in let sigma = Tacmach.New.project gl in Coqlib.check_required_library ["Coq";"fourier";"Fourier"]; - let goal = Termops.strip_outer_cast sigma (EConstr.of_constr concl) in + let goal = Termops.strip_outer_cast sigma concl in let goal = EConstr.Unsafe.to_constr goal in let fhyp=Id.of_string "new_hyp_for_fourier" in (* si le but est une inéquation, on introduit son contraire, diff --git a/plugins/omega/coq_omega.ml b/plugins/omega/coq_omega.ml index f39549514..72290e681 100644 --- a/plugins/omega/coq_omega.ml +++ b/plugins/omega/coq_omega.ml @@ -1857,6 +1857,7 @@ let destructure_goal = in Tacticals.New.tclTHEN goal_tac destructure_hyps in + let concl = EConstr.Unsafe.to_constr concl in (loop concl) end } diff --git a/plugins/quote/quote.ml b/plugins/quote/quote.ml index 04a747fb4..5f8a0b2d5 100644 --- a/plugins/quote/quote.ml +++ b/plugins/quote/quote.ml @@ -452,6 +452,7 @@ let quote f lid = let cl = List.map (fun id -> Tacmach.New.pf_global id gl) lid in let ivs = compute_ivs f cl gl in let concl = Proofview.Goal.concl gl in + let concl = EConstr.Unsafe.to_constr concl in let quoted_terms = quote_terms ivs [concl] in let (p, vm) = match quoted_terms with | [p], vm -> (p,vm) diff --git a/plugins/ssrmatching/ssrmatching.ml4 b/plugins/ssrmatching/ssrmatching.ml4 index aa91494eb..d34c9325e 100644 --- a/plugins/ssrmatching/ssrmatching.ml4 +++ b/plugins/ssrmatching/ssrmatching.ml4 @@ -1078,7 +1078,7 @@ END let thin id sigma goal = let ids = Id.Set.singleton id in let env = Goal.V82.env sigma goal in - let cl = EConstr.Unsafe.to_constr (Goal.V82.concl sigma goal) in + let cl = Goal.V82.concl sigma goal in let evdref = ref (Evd.clear_metas sigma) in let ans = try Some (Evarutil.clear_hyps_in_evi env evdref (Environ.named_context_val env) cl ids) diff --git a/proofs/clenvtac.ml b/proofs/clenvtac.ml index 84cdc09ee..32c887ff5 100644 --- a/proofs/clenvtac.ml +++ b/proofs/clenvtac.ml @@ -144,7 +144,7 @@ let unify ?(flags=fail_quick_unif_flags) m = let n = Tacmach.New.pf_concl (Proofview.Goal.assume gl) in let evd = clear_metas (Tacmach.New.project gl) in try - let evd' = w_unify env evd CONV ~flags m (EConstr.of_constr n) in + let evd' = w_unify env evd CONV ~flags m n in Proofview.Unsafe.tclEVARSADVANCE evd' with e when CErrors.noncritical e -> Proofview.tclZERO e end } diff --git a/proofs/refine.ml b/proofs/refine.ml index 11eabe0a9..32064aff5 100644 --- a/proofs/refine.ml +++ b/proofs/refine.ml @@ -77,7 +77,6 @@ let make_refine_enter ?(unsafe = true) f = let sigma = Sigma.to_evar_map sigma in let env = Proofview.Goal.env gl in let concl = Proofview.Goal.concl gl in - let concl = EConstr.of_constr concl in (** Save the [future_goals] state to restore them after the refinement. *) let prev_future_goals = Evd.future_goals sigma in @@ -146,7 +145,6 @@ let with_type env evd c t = let refine_casted ?unsafe f = Proofview.Goal.enter { enter = begin fun gl -> let gl = Proofview.Goal.assume gl in let concl = Proofview.Goal.concl gl in - let concl = EConstr.of_constr concl in let env = Proofview.Goal.env gl in let f = { run = fun h -> let Sigma (c, h, p) = f.run h in diff --git a/proofs/tacmach.ml b/proofs/tacmach.ml index aa54e4f9b..f3f19e854 100644 --- a/proofs/tacmach.ml +++ b/proofs/tacmach.ml @@ -220,7 +220,7 @@ module New = struct let gl = Proofview.Goal.assume gl in let concl = Proofview.Goal.concl gl in let sigma = project gl in - nf_evar sigma concl + EConstr.of_constr (nf_evar sigma (EConstr.Unsafe.to_constr concl)) let pf_whd_all gl t = pf_apply whd_all gl t @@ -229,13 +229,13 @@ module New = struct let pf_reduce_to_quantified_ind gl t = pf_apply reduce_to_quantified_ind gl t - let pf_hnf_constr gl t = pf_apply hnf_constr gl t + let pf_hnf_constr gl t = EConstr.of_constr (pf_apply hnf_constr gl t) let pf_hnf_type_of gl t = - pf_whd_all gl (EConstr.of_constr (pf_get_type_of gl t)) + EConstr.of_constr (pf_whd_all gl (EConstr.of_constr (pf_get_type_of gl t))) let pf_matches gl pat t = pf_apply Constr_matching.matches_conv gl pat t - let pf_whd_all gl t = pf_apply whd_all gl t + let pf_whd_all gl t = EConstr.of_constr (pf_apply whd_all gl t) let pf_compute gl t = pf_apply compute gl t let pf_nf_evar gl t = nf_evar (project gl) t diff --git a/proofs/tacmach.mli b/proofs/tacmach.mli index 340c7addc..7b387ac9a 100644 --- a/proofs/tacmach.mli +++ b/proofs/tacmach.mli @@ -106,7 +106,7 @@ module New : sig val project : ('a, 'r) Proofview.Goal.t -> Evd.evar_map val pf_env : ('a, 'r) Proofview.Goal.t -> Environ.env - val pf_concl : ([ `NF ], 'r) Proofview.Goal.t -> types + val pf_concl : ([ `NF ], 'r) Proofview.Goal.t -> EConstr.types val pf_unsafe_type_of : ('a, 'r) Proofview.Goal.t -> EConstr.constr -> Term.types val pf_type_of : ('a, 'r) Proofview.Goal.t -> EConstr.constr -> evar_map * Term.types @@ -120,13 +120,13 @@ module New : sig val pf_get_hyp_typ : identifier -> ([ `NF ], 'r) Proofview.Goal.t -> types val pf_last_hyp : ([ `NF ], 'r) Proofview.Goal.t -> Context.Named.Declaration.t - val pf_nf_concl : ([ `LZ ], 'r) Proofview.Goal.t -> types + val pf_nf_concl : ([ `LZ ], 'r) Proofview.Goal.t -> EConstr.types val pf_reduce_to_quantified_ind : ('a, 'r) Proofview.Goal.t -> EConstr.types -> pinductive * EConstr.types - val pf_hnf_constr : ('a, 'r) Proofview.Goal.t -> EConstr.constr -> types - val pf_hnf_type_of : ('a, 'r) Proofview.Goal.t -> EConstr.constr -> types + val pf_hnf_constr : ('a, 'r) Proofview.Goal.t -> EConstr.constr -> EConstr.types + val pf_hnf_type_of : ('a, 'r) Proofview.Goal.t -> EConstr.constr -> EConstr.types - val pf_whd_all : ('a, 'r) Proofview.Goal.t -> EConstr.t -> constr + val pf_whd_all : ('a, 'r) Proofview.Goal.t -> EConstr.t -> EConstr.constr val pf_compute : ('a, 'r) Proofview.Goal.t -> EConstr.t -> EConstr.constr val pf_matches : ('a, 'r) Proofview.Goal.t -> constr_pattern -> EConstr.constr -> patvar_map diff --git a/tactics/auto.ml b/tactics/auto.ml index 21fe9667b..4218be0bb 100644 --- a/tactics/auto.ml +++ b/tactics/auto.ml @@ -331,7 +331,6 @@ let rec trivial_fail_db dbg mod_delta db_list local_db = in Proofview.Goal.enter { enter = begin fun gl -> let concl = Tacmach.New.pf_nf_concl gl in - let concl = EConstr.of_constr concl in let sigma = Tacmach.New.project gl in let secvars = compute_secvars gl in Tacticals.New.tclFIRST @@ -492,7 +491,6 @@ let search d n mod_delta db_list local_db = (Tacticals.New.tclORELSE0 (intro_register d (search d n) local_db) ( Proofview.Goal.enter { enter = begin fun gl -> let concl = Tacmach.New.pf_nf_concl gl in - let concl = EConstr.of_constr concl in let sigma = Tacmach.New.project gl in let secvars = compute_secvars gl in let d' = incr_dbg d in diff --git a/tactics/class_tactics.ml b/tactics/class_tactics.ml index 6c45bef1c..bf4e53b10 100644 --- a/tactics/class_tactics.ml +++ b/tactics/class_tactics.ml @@ -249,7 +249,6 @@ let unify_resolve_refine poly flags = { enter = begin fun gls ((c, t, ctx),n,clenv) -> let env = Proofview.Goal.env gls in let concl = Proofview.Goal.concl gls in - let concl = EConstr.of_constr concl in Refine.refine ~unsafe:true { Sigma.run = fun sigma -> let sigma = Sigma.to_evar_map sigma in let sigma, term, ty = @@ -363,7 +362,7 @@ let rec e_trivial_fail_db only_classes db_list local_db secvars = Proofview.Goal.nf_enter { enter = begin fun gl -> let tacs = e_trivial_resolve db_list local_db secvars only_classes - (project gl) (EConstr.of_constr (pf_concl gl)) in + (project gl) (pf_concl gl) in tclFIRST (List.map (fun (x,_,_,_,_) -> x) tacs) end} in @@ -1002,7 +1001,6 @@ module Search = struct let open Proofview.Notations in let env = Goal.env gl in let concl = Goal.concl gl in - let concl = EConstr.of_constr concl in let sigma = Goal.sigma gl in let s = Sigma.to_evar_map sigma in let unique = not info.search_dep || is_unique env s concl in diff --git a/tactics/eauto.ml b/tactics/eauto.ml index 986d1a624..57400d334 100644 --- a/tactics/eauto.ml +++ b/tactics/eauto.ml @@ -34,7 +34,6 @@ let e_give_exact ?(flags=eauto_unif_flags) c = let t1 = Tacmach.New.pf_unsafe_type_of gl c in let t1 = EConstr.of_constr t1 in let t2 = Tacmach.New.pf_concl (Proofview.Goal.assume gl) in - let t2 = EConstr.of_constr t2 in let sigma = Tacmach.New.project gl in if occur_existential sigma t1 || occur_existential sigma t2 then Tacticals.New.tclTHEN (Clenvtac.unify ~flags t1) (exact_no_check c) @@ -151,7 +150,7 @@ let rec e_trivial_fail_db db_list local_db = let tacl = registered_e_assumption :: (Tacticals.New.tclTHEN Tactics.intro next) :: - (List.map fst (e_trivial_resolve (Tacmach.New.project gl) db_list local_db secvars (EConstr.of_constr (Tacmach.New.pf_nf_concl gl)))) + (List.map fst (e_trivial_resolve (Tacmach.New.project gl) db_list local_db secvars (Tacmach.New.pf_nf_concl gl))) in Tacticals.New.tclFIRST (List.map Tacticals.New.tclCOMPLETE tacl) end } @@ -508,7 +507,6 @@ let autounfold_one db cl = let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in let concl = Proofview.Goal.concl gl in - let concl = EConstr.of_constr concl in let st = List.fold_left (fun (i,c) dbname -> let db = try searchtable_map dbname diff --git a/tactics/eqdecide.ml b/tactics/eqdecide.ml index a96656d3a..16e0d9684 100644 --- a/tactics/eqdecide.ml +++ b/tactics/eqdecide.ml @@ -178,7 +178,6 @@ let solveEqBranch rectype = begin Proofview.Goal.enter { enter = begin fun gl -> let concl = pf_nf_concl gl in - let concl = EConstr.of_constr concl in let sigma = project gl in match_eqdec sigma concl >>= fun (eqonleft,op,lhs,rhs,_) -> let (mib,mip) = Global.lookup_inductive rectype in @@ -205,7 +204,6 @@ let decideGralEquality = begin Proofview.Goal.enter { enter = begin fun gl -> let concl = pf_nf_concl gl in - let concl = EConstr.of_constr concl in let sigma = project gl in match_eqdec sigma concl >>= fun (eqonleft,_,c1,c2,typ) -> let headtyp = hd_app sigma (pf_compute gl typ) in diff --git a/tactics/equality.ml b/tactics/equality.ml index 2eead2d22..209c9eb66 100644 --- a/tactics/equality.ml +++ b/tactics/equality.ml @@ -312,9 +312,8 @@ let general_elim_clause with_evars frzevars tac cls c t l l2r elim = in let typ = match cls with | None -> pf_nf_concl gl - | Some id -> pf_get_hyp_typ id (Proofview.Goal.assume gl) + | Some id -> EConstr.of_constr (pf_get_hyp_typ id (Proofview.Goal.assume gl)) in - let typ = EConstr.of_constr typ in let cs = instantiate_lemma typ in if firstonly then tclFIRST (List.map try_clause cs) else tclMAP try_clause cs @@ -409,7 +408,7 @@ let find_elim hdcncl lft2rgt dep cls ot gl = let type_of_clause cls gl = match cls with | None -> Proofview.Goal.concl gl - | Some id -> pf_get_hyp_typ id gl + | Some id -> EConstr.of_constr (pf_get_hyp_typ id gl) let leibniz_rewrite_ebindings_clause cls lft2rgt tac c t l with_evars frzevars dep_proof_ok hdcncl = Proofview.Goal.nf_s_enter { s_enter = begin fun gl -> @@ -417,7 +416,6 @@ let leibniz_rewrite_ebindings_clause cls lft2rgt tac c t l with_evars frzevars d let isatomic = isProd evd (EConstr.of_constr (whd_zeta evd hdcncl)) in let dep_fun = if isatomic then dependent else dependent_no_evar in let type_of_cls = type_of_clause cls gl in - let type_of_cls = EConstr.of_constr type_of_cls in let dep = dep_proof_ok && dep_fun evd c type_of_cls in let Sigma ((elim, effs), sigma, p) = find_elim hdcncl lft2rgt dep cls (Some t) gl in let tac = @@ -1052,7 +1050,6 @@ let onNegatedEquality with_evars tac = Proofview.Goal.nf_enter { enter = begin fun gl -> let sigma = Tacmach.New.project gl in let ccl = Proofview.Goal.concl gl in - let ccl = EConstr.of_constr ccl in let env = Proofview.Goal.env gl in match EConstr.kind sigma (EConstr.of_constr (hnf_constr env sigma ccl)) with | Prod (_,t,u) when is_empty_type sigma u -> @@ -1611,7 +1608,6 @@ let cutSubstInConcl l2r eqn = let sigma = Proofview.Goal.sigma gl in let (lbeq,u,(t,e1,e2)) = find_eq_data_decompose gl eqn in let typ = pf_concl gl in - let typ = EConstr.of_constr typ in let (e1,e2) = if l2r then (e1,e2) else (e2,e1) in let Sigma ((typ, expected), sigma, p) = subst_tuple_term env sigma e1 e2 typ in let tac = @@ -1752,7 +1748,7 @@ let subst_one dep_proof_ok x (hyp,rhs,dir) = hyps (MoveBefore x,[x],[]))) in (* In practice, no dep hyps before x, so MoveBefore x is good enough *) (* Decides if x appears in conclusion *) - let depconcl = occur_var env sigma x (EConstr.of_constr concl) in + let depconcl = occur_var env sigma x concl in let need_rewrite = not (List.is_empty dephyps) || depconcl in tclTHENLIST ((if need_rewrite then diff --git a/tactics/hipattern.ml b/tactics/hipattern.ml index b92d65908..fa114a3d3 100644 --- a/tactics/hipattern.ml +++ b/tactics/hipattern.ml @@ -468,7 +468,6 @@ let match_eq_nf gls eqn (ref, hetero) = let n = if hetero then 4 else 3 in let args = List.init n (fun i -> mkGPatVar ("X" ^ string_of_int (i + 1))) in let pat = mkPattern (mkGAppRef ref args) in - let pf_whd_all gls c = EConstr.of_constr (pf_whd_all gls c) in match Id.Map.bindings (pf_matches gls pat eqn) with | [(m1,t);(m2,x);(m3,y)] -> assert (Id.equal m1 meta1 && Id.equal m2 meta2 && Id.equal m3 meta3); diff --git a/tactics/inv.ml b/tactics/inv.ml index ac9a564e5..e45eb2a16 100644 --- a/tactics/inv.ml +++ b/tactics/inv.ml @@ -34,7 +34,7 @@ module NamedDecl = Context.Named.Declaration let var_occurs_in_pf gl id = let env = Proofview.Goal.env gl in let sigma = project gl in - occur_var env sigma id (EConstr.of_constr (Proofview.Goal.concl gl)) || + occur_var env sigma id (Proofview.Goal.concl gl) || List.exists (occur_var_in_decl env sigma id) (Proofview.Goal.hyps gl) (* [make_inv_predicate (ity,args) C] @@ -441,7 +441,6 @@ let raw_inversion inv_kind id status names = let sigma = Sigma.to_evar_map sigma in let env = Proofview.Goal.env gl in let concl = Proofview.Goal.concl gl in - let concl = EConstr.of_constr concl in let c = mkVar id in let (ind, t) = try pf_apply Tacred.reduce_to_atomic_ind gl (EConstr.of_constr (pf_unsafe_type_of gl c)) @@ -522,13 +521,13 @@ let invIn k names ids id = let hyps = List.map (fun id -> pf_get_hyp id gl) ids in let concl = Proofview.Goal.concl gl in let sigma = project gl in - let nb_prod_init = nb_prod sigma (EConstr.of_constr concl) in + let nb_prod_init = nb_prod sigma concl in let intros_replace_ids = Proofview.Goal.enter { enter = begin fun gl -> let concl = pf_nf_concl gl in let sigma = project gl in let nb_of_new_hyp = - nb_prod sigma (EConstr.of_constr concl) - (List.length hyps + nb_prod_init) + nb_prod sigma concl - (List.length hyps + nb_prod_init) in if nb_of_new_hyp < 1 then intros_replacing ids diff --git a/tactics/leminv.ml b/tactics/leminv.ml index 62e78b588..609fa1be8 100644 --- a/tactics/leminv.ml +++ b/tactics/leminv.ml @@ -294,7 +294,7 @@ let lemInvIn id c ids = let intros_replace_ids = let concl = Proofview.Goal.concl gl in let sigma = project gl in - let nb_of_new_hyp = nb_prod sigma (EConstr.of_constr concl) - List.length ids in + let nb_of_new_hyp = nb_prod sigma concl - List.length ids in if nb_of_new_hyp < 1 then intros_replacing ids else diff --git a/tactics/tacticals.ml b/tactics/tacticals.ml index e440f1dc5..d79a74b36 100644 --- a/tactics/tacticals.ml +++ b/tactics/tacticals.ml @@ -710,7 +710,7 @@ module New = struct let elimination_sort_of_goal gl = (** Retyping will expand evars anyway. *) let c = Proofview.Goal.concl (Goal.assume gl) in - pf_apply Retyping.get_sort_family_of gl (EConstr.of_constr c) + pf_apply Retyping.get_sort_family_of gl c let elimination_sort_of_hyp id gl = (** Retyping will expand evars anyway. *) diff --git a/tactics/tactics.ml b/tactics/tactics.ml index 03f81773b..dabe78b34 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -207,7 +207,6 @@ let introduction ?(check=true) id = Proofview.Goal.enter { enter = begin fun gl -> let gl = Proofview.Goal.assume gl in let concl = Proofview.Goal.concl gl in - let concl = EConstr.of_constr concl in let sigma = Tacmach.New.project gl in let hyps = named_context_val (Proofview.Goal.env gl) in let store = Proofview.Goal.extra gl in @@ -230,7 +229,6 @@ let convert_concl ?(check=true) ty k = let env = Proofview.Goal.env gl in let store = Proofview.Goal.extra gl in let conclty = Proofview.Goal.raw_concl gl in - let conclty = EConstr.of_constr conclty in Refine.refine ~unsafe:true { run = begin fun sigma -> let Sigma ((), sigma, p) = if check then begin @@ -251,7 +249,6 @@ let convert_hyp ?(check=true) d = let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in let ty = Proofview.Goal.raw_concl gl in - let ty = EConstr.of_constr ty in let store = Proofview.Goal.extra gl in let sign = convert_hyp check (named_context_val env) sigma d in let env = reset_with_named_context sign env in @@ -353,7 +350,6 @@ let move_hyp id dest = let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in let ty = Proofview.Goal.raw_concl gl in - let ty = EConstr.of_constr ty in let store = Proofview.Goal.extra gl in let sign = named_context_val env in let sign' = move_hyp_in_named_context sigma id dest sign in @@ -409,6 +405,7 @@ let rename_hyp repl = |> NamedDecl.map_constr subst in let nhyps = List.map map hyps in + let concl = EConstr.Unsafe.to_constr concl in let nconcl = subst concl in let nconcl = EConstr.of_constr nconcl in let nctx = Environ.val_of_named_context nhyps in @@ -545,7 +542,6 @@ let mutual_fix f n rest j = Proofview.Goal.nf_enter { enter = begin fun gl -> let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in let concl = Proofview.Goal.concl gl in - let concl = EConstr.of_constr concl in let (sp, u) = check_mutind env sigma n concl in let firsts, lasts = List.chop j rest in let all = firsts @ (f, n, concl) :: lasts in @@ -601,7 +597,6 @@ let mutual_cofix f others j = Proofview.Goal.nf_enter { enter = begin fun gl -> let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in let concl = Proofview.Goal.concl gl in - let concl = EConstr.of_constr concl in let firsts,lasts = List.chop j others in let all = firsts @ (f, concl) :: lasts in List.iter (fun (_, c) -> check_is_mutcoind env sigma c) all; @@ -727,7 +722,7 @@ let bind_red_expr_occurrences occs nbcl redexp = let reduct_in_concl (redfun,sty) = Proofview.Goal.nf_enter { enter = begin fun gl -> - convert_concl_no_check (EConstr.of_constr (Tacmach.New.pf_apply redfun gl (EConstr.of_constr (Tacmach.New.pf_concl gl)))) sty + convert_concl_no_check (EConstr.of_constr (Tacmach.New.pf_apply redfun gl (Tacmach.New.pf_concl gl))) sty end } let reduct_in_hyp ?(check=false) redfun (id,where) = @@ -762,7 +757,7 @@ let pf_e_reduce_decl redfun where decl gl = let e_reduct_in_concl ~check (redfun, sty) = Proofview.Goal.nf_s_enter { s_enter = begin fun gl -> let sigma = Proofview.Goal.sigma gl in - let Sigma (c', sigma, p) = redfun.e_redfun (Tacmach.New.pf_env gl) sigma (EConstr.of_constr (Tacmach.New.pf_concl gl)) in + let Sigma (c', sigma, p) = redfun.e_redfun (Tacmach.New.pf_env gl) sigma (Tacmach.New.pf_concl gl) in let c' = EConstr.of_constr c' in Sigma (convert_concl ~check c' sty, sigma, p) end } @@ -783,7 +778,7 @@ let e_reduct_option ?(check=false) redfun = function let e_change_in_concl (redfun,sty) = Proofview.Goal.s_enter { s_enter = begin fun gl -> let sigma = Proofview.Goal.sigma gl in - let Sigma (c, sigma, p) = redfun.e_redfun (Proofview.Goal.env gl) sigma (EConstr.of_constr (Proofview.Goal.raw_concl gl)) in + let Sigma (c, sigma, p) = redfun.e_redfun (Proofview.Goal.env gl) sigma (Proofview.Goal.raw_concl gl) in let c = EConstr.of_constr c in Sigma (convert_concl_no_check c sty, sigma, p) end } @@ -976,7 +971,6 @@ let rec intro_then_gen name_flag move_flag force_flag dep_flag tac = Proofview.Goal.enter { enter = begin fun gl -> let sigma = Tacmach.New.project gl in let concl = Proofview.Goal.concl (Proofview.Goal.assume gl) in - let concl = EConstr.of_constr concl in match EConstr.kind sigma concl with | Prod (name,t,u) when not dep_flag || not (noccurn sigma 1 u) -> let name = find_name false (local_assum (name,t)) name_flag gl in @@ -1120,7 +1114,7 @@ let lookup_hypothesis_as_renamed_gen red h gl = aux c | x -> x in - try aux (Proofview.Goal.concl gl) + try aux (EConstr.to_constr (Tacmach.New.project gl) (Proofview.Goal.concl gl)) with Redelimination -> None let is_quantified_hypothesis id gl = @@ -1261,7 +1255,6 @@ let cut c = let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in let concl = Tacmach.New.pf_nf_concl gl in - let concl = EConstr.of_constr concl in let is_sort = try (** Backward compat: ensure that [c] is well-typed. *) @@ -1302,7 +1295,7 @@ let check_unresolved_evars_of_metas sigma clenv = (match kind_of_term c.rebus with | Evar (evk,_) when Evd.is_undefined clenv.evd evk && not (Evd.mem sigma evk) -> - error_uninstantiated_metas (EConstr.mkMeta mv) clenv + error_uninstantiated_metas (mkMeta mv) clenv | _ -> ()) | _ -> ()) (meta_list clenv.evd) @@ -1401,7 +1394,6 @@ let enforce_prop_bound_names rename tac = let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in let t = Proofview.Goal.concl gl in - let t = EConstr.of_constr t in change_concl (aux env sigma i t) end } in (if isrec then Tacticals.New.tclTHENFIRSTn else Tacticals.New.tclTHENLASTn) @@ -1487,7 +1479,7 @@ let general_case_analysis_in_context with_evars clear_flag (c,lbindc) = let (mind,_) = reduce_to_quantified_ind env (Sigma.to_evar_map sigma) t in let sort = Tacticals.New.elimination_sort_of_goal gl in let Sigma (elim, sigma, p) = - if occur_term (Sigma.to_evar_map sigma) c (EConstr.of_constr concl) then + if occur_term (Sigma.to_evar_map sigma) c concl then build_case_analysis_scheme env sigma mind true sort else build_case_analysis_scheme_default env sigma mind sort in @@ -1728,7 +1720,6 @@ let solve_remaining_apply_goals = let env = Proofview.Goal.env gl in let evd = Sigma.to_evar_map sigma in let concl = Proofview.Goal.concl gl in - let concl = EConstr.of_constr concl in if Typeclasses.is_class_type evd concl then let evd', c' = Typeclasses.resolve_one_typeclass env evd concl in let tac = Refine.refine ~unsafe:true { run = fun h -> Sigma.here c' h } in @@ -1755,7 +1746,7 @@ let general_apply with_delta with_destruct with_evars clear_flag (loc,(c,lbind : (* The actual type of the theorem. It will be matched against the goal. If this fails, then the head constant will be unfolded step by step. *) - let concl_nprod = nb_prod_modulo_zeta sigma (EConstr.of_constr concl) in + let concl_nprod = nb_prod_modulo_zeta sigma concl in let rec try_main_apply with_destruct c = Proofview.Goal.enter { enter = begin fun gl -> let env = Proofview.Goal.env gl in @@ -1966,10 +1957,9 @@ let apply_in_delayed_once sidecond_first with_delta with_destruct with_evars nam let cut_and_apply c = Proofview.Goal.nf_enter { enter = begin fun gl -> let sigma = Tacmach.New.project gl in - match EConstr.kind sigma (EConstr.of_constr (Tacmach.New.pf_hnf_constr gl (EConstr.of_constr (Tacmach.New.pf_unsafe_type_of gl c)))) with + match EConstr.kind sigma (Tacmach.New.pf_hnf_constr gl (EConstr.of_constr (Tacmach.New.pf_unsafe_type_of gl c))) with | Prod (_,c1,c2) when Vars.noccurn sigma 1 c2 -> let concl = Proofview.Goal.concl gl in - let concl = EConstr.of_constr concl in let env = Tacmach.New.pf_env gl in Refine.refine { run = begin fun sigma -> let typ = mkProd (Anonymous, c2, concl) in @@ -1999,7 +1989,6 @@ let exact_check c = let sigma = Proofview.Goal.sigma gl in (** We do not need to normalize the goal because we just check convertibility *) let concl = Proofview.Goal.concl (Proofview.Goal.assume gl) in - let concl = EConstr.of_constr concl in let env = Proofview.Goal.env gl in let sigma = Sigma.to_evar_map sigma in let sigma, ct = Typing.type_of env sigma c in @@ -2013,8 +2002,7 @@ let exact_check c = let cast_no_check cast c = Proofview.Goal.enter { enter = begin fun gl -> let concl = Proofview.Goal.concl (Proofview.Goal.assume gl) in - let concl = EConstr.of_constr concl in - exact_no_check (EConstr.mkCast (c, cast, concl)) + exact_no_check (mkCast (c, cast, concl)) end } let vm_cast_no_check c = cast_no_check Term.VMcast c @@ -2025,7 +2013,7 @@ let exact_proof c = Proofview.Goal.nf_enter { enter = begin fun gl -> Refine.refine { run = begin fun sigma -> let sigma = Sigma.to_evar_map sigma in - let (c, ctx) = Constrintern.interp_casted_constr (pf_env gl) sigma c (pf_concl gl) in + let (c, ctx) = Constrintern.interp_casted_constr (pf_env gl) sigma c (EConstr.Unsafe.to_constr (pf_concl gl)) in let c = EConstr.of_constr c in let sigma = Evd.merge_universe_context sigma ctx in Sigma.Unsafe.of_pair (c, sigma) @@ -2041,13 +2029,14 @@ let assumption = else Tacticals.New.tclZEROMSG (str "No such assumption.") | decl::rest -> let t = NamedDecl.get_type decl in + let t = EConstr.of_constr t in let concl = Proofview.Goal.concl gl in let sigma = Tacmach.New.project gl in let (sigma, is_same_type) = - if only_eq then (sigma, Constr.equal t concl) + if only_eq then (sigma, EConstr.eq_constr sigma t concl) else let env = Proofview.Goal.env gl in - infer_conv env sigma (EConstr.of_constr t) (EConstr.of_constr concl) + infer_conv env sigma t concl in if is_same_type then (Proofview.Unsafe.tclEVARS sigma) <*> @@ -2099,7 +2088,6 @@ let clear_body ids = Proofview.Goal.enter { enter = begin fun gl -> let env = Proofview.Goal.env gl in let concl = Proofview.Goal.concl (Proofview.Goal.assume gl) in - let concl = EConstr.of_constr concl in let sigma = Tacmach.New.project gl in let ctx = named_context env in let map = function @@ -2173,7 +2161,7 @@ let keep hyps = let hyp = NamedDecl.get_id decl in if Id.List.mem hyp hyps || List.exists (occur_var_in_decl env sigma hyp) keep - || occur_var env sigma hyp (EConstr.of_constr ccl) + || occur_var env sigma hyp ccl then (clear,decl::keep) else (hyp::clear,keep)) ~init:([],[]) (Proofview.Goal.env gl) @@ -2213,7 +2201,6 @@ let bring_hyps hyps = let env = Proofview.Goal.env gl in let store = Proofview.Goal.extra gl in let concl = Tacmach.New.pf_nf_concl gl in - let concl = EConstr.of_constr concl in let newcl = List.fold_right mkNamedProd_or_LetIn hyps concl in let args = Array.map_of_list EConstr.of_constr (Context.Named.to_instance hyps) in Refine.refine { run = begin fun sigma -> @@ -2251,7 +2238,6 @@ let constructor_tac with_evars expctdnumopt i lbind = let reduce_to_quantified_ind = Tacmach.New.pf_apply Tacred.reduce_to_quantified_ind gl in - let cl = EConstr.of_constr cl in let (mind,redcl) = reduce_to_quantified_ind cl in let nconstr = Array.length (snd (Global.lookup_inductive (fst mind))).mind_consnames in @@ -2291,7 +2277,6 @@ let any_constructor with_evars tacopt = let reduce_to_quantified_ind = Tacmach.New.pf_apply Tacred.reduce_to_quantified_ind gl in - let cl = EConstr.of_constr cl in let mind = fst (reduce_to_quantified_ind cl) in let nconstr = Array.length (snd (Global.lookup_inductive (fst mind))).mind_consnames in @@ -2768,7 +2753,6 @@ let letin_tac with_eq id c ty occs = let env = Proofview.Goal.env gl in let ccl = Proofview.Goal.concl gl in let abs = AbstractExact (id,c,ty,occs,true) in - let ccl = EConstr.of_constr ccl in let (id,_,depdecls,lastlhyp,ccl,res) = make_abstraction env sigma ccl abs in (* We keep the original term to match but record the potential side-effects of unifying universes. *) @@ -2787,7 +2771,6 @@ let letin_pat_tac with_eq id c occs = let ccl = Proofview.Goal.concl gl in let check t = true in let abs = AbstractPattern (false,check,id,c,occs,false) in - let ccl = EConstr.of_constr ccl in let (id,_,depdecls,lastlhyp,ccl,res) = make_abstraction env sigma ccl abs in let Sigma (c, sigma, p) = match res with | None -> finish_evar_resolution ~flags:(tactic_infer_flags false) env sigma c @@ -2921,7 +2904,7 @@ let generalize_gen_let lconstr = Proofview.Goal.nf_s_enter { s_enter = begin fun let env = Proofview.Goal.env gl in let newcl, evd = List.fold_right_i (Tacmach.New.of_old generalize_goal gl) 0 lconstr - (EConstr.of_constr (Tacmach.New.pf_concl gl),Tacmach.New.project gl) + (Tacmach.New.pf_concl gl,Tacmach.New.project gl) in let (evd, _) = Typing.type_of env evd newcl in let map ((_, c, b),_) = if Option.is_empty b then Some c else None in @@ -2934,7 +2917,6 @@ let new_generalize_gen_let lconstr = let sigma = Proofview.Goal.sigma gl in let gl = Proofview.Goal.assume gl in let concl = Proofview.Goal.concl gl in - let concl = EConstr.of_constr concl in let sigma = Sigma.to_evar_map sigma in let env = Proofview.Goal.env gl in let ids = Tacmach.New.pf_ids_of_hyps gl in @@ -3475,8 +3457,8 @@ let cook_sign hyp0_opt inhyps indvars env sigma = (* [rel_contexts] and [rel_declaration] actually contain triples, and lists are actually in reverse order to fit [compose_prod]. *) type elim_scheme = { - elimc: EConstr.constr with_bindings option; - elimt: EConstr.types; + elimc: constr with_bindings option; + elimt: types; indref: global_reference option; params: Context.Rel.t; (* (prm1,tprm1);(prm2,tprm2)...(prmp,tprmp) *) nparams: int; (* number of parameters *) @@ -3488,7 +3470,7 @@ type elim_scheme = { nargs: int; (* number of arguments *) indarg: Context.Rel.Declaration.t option; (* Some (H,I prm1..prmp x1...xni) if HI is in premisses, None otherwise *) - concl: EConstr.types; (* Qi x1...xni HI (f...), HI and (f...) + concl: types; (* Qi x1...xni HI (f...), HI and (f...) are optional and mutually exclusive *) indarg_in_concl: bool; (* true if HI appears at the end of conclusion *) farg_in_concl: bool; (* true if (f...) appears at the end of conclusion *) @@ -3497,7 +3479,7 @@ type elim_scheme = { let empty_scheme = { elimc = None; - elimt = EConstr.mkProp; + elimt = mkProp; indref = None; params = []; nparams = 0; @@ -3508,7 +3490,7 @@ let empty_scheme = args = []; nargs = 0; indarg = None; - concl = EConstr.mkProp; + concl = mkProp; indarg_in_concl = false; farg_in_concl = false; } @@ -3582,7 +3564,7 @@ let lift_togethern n l = l ([], n) in l' -let lift_list l = List.map (EConstr.Vars.lift 1) l +let lift_list l = List.map (lift 1) l let ids_of_constr sigma ?(all=false) vars c = let rec aux vars c = @@ -4251,7 +4233,6 @@ let apply_induction_in_context with_evars hyp0 inhyps elim indvars names induct_ let env = Proofview.Goal.env gl in let sigma = Sigma.to_evar_map sigma in let concl = Tacmach.New.pf_nf_concl gl in - let concl = EConstr.of_constr concl in let statuslists,lhyp0,toclear,deps,avoid,dep_in_hyps = cook_sign hyp0 inhyps indvars env sigma in let dep_in_concl = Option.cata (fun id -> occur_var env sigma id concl) false hyp0 in let dep = dep_in_hyps || dep_in_concl in @@ -4341,7 +4322,7 @@ let induction_without_atomization isrec with_evars elim names lid = (* assume that no occurrences are selected *) let clear_unselected_context id inhyps cls = Proofview.Goal.nf_enter { enter = begin fun gl -> - if occur_var (Tacmach.New.pf_env gl) (Tacmach.New.project gl) id (EConstr.of_constr (Tacmach.New.pf_concl gl)) && + if occur_var (Tacmach.New.pf_env gl) (Tacmach.New.project gl) id (Tacmach.New.pf_concl gl) && cls.concl_occs == NoOccurrences then user_err (str "Conclusion must be mentioned: it depends on " ++ pr_id id @@ -4437,7 +4418,6 @@ let pose_induction_arg_then isrec with_evars (is_arg_pure_hyp,from_prefix) elim let check = check_enough_applied env sigma elim in let Sigma (c, sigma', p) = use_bindings env sigma elim false (c0,lbind) t0 in let abs = AbstractPattern (from_prefix,check,Name id,(pending,c),cls,false) in - let ccl = EConstr.of_constr ccl in let (id,sign,_,lastlhyp,ccl,res) = make_abstraction env sigma' ccl abs in match res with | None -> @@ -4492,7 +4472,7 @@ let pose_induction_arg_then isrec with_evars (is_arg_pure_hyp,from_prefix) elim let has_generic_occurrences_but_goal cls id env sigma ccl = clause_with_generic_context_selection cls && (* TODO: whd_evar of goal *) - (cls.concl_occs != NoOccurrences || not (occur_var env sigma id (EConstr.of_constr ccl))) + (cls.concl_occs != NoOccurrences || not (occur_var env sigma id ccl)) let induction_gen clear_flag isrec with_evars elim ((_pending,(c,lbind)),(eqname,names) as arg) cls = @@ -4736,7 +4716,7 @@ let maybe_betadeltaiota_concl allowred gl = if not allowred then concl else let env = Proofview.Goal.env gl in - whd_all env sigma (EConstr.of_constr concl) + EConstr.of_constr (whd_all env sigma concl) let reflexivity_red allowred = Proofview.Goal.enter { enter = begin fun gl -> @@ -4745,7 +4725,6 @@ let reflexivity_red allowred = inside setoid_reflexivity (see Optimize cases in setoid_replace.ml). *) let sigma = Tacmach.New.project gl in let concl = maybe_betadeltaiota_concl allowred gl in - let concl = EConstr.of_constr concl in match match_with_equality_type sigma concl with | None -> Proofview.tclZERO NoEquationFound | Some _ -> one_constructor 1 NoBindings @@ -4797,7 +4776,6 @@ let symmetry_red allowred = inside setoid_reflexivity (see Optimize cases in setoid_replace.ml). *) let sigma = Tacmach.New.project gl in let concl = maybe_betadeltaiota_concl allowred gl in - let concl = EConstr.of_constr concl in match_with_equation sigma concl >>= fun with_eqn -> match with_eqn with | Some eq_data,_,_ -> @@ -4895,7 +4873,6 @@ let transitivity_red allowred t = inside setoid_reflexivity (see Optimize cases in setoid_replace.ml). *) let sigma = Tacmach.New.project gl in let concl = maybe_betadeltaiota_concl allowred gl in - let concl = EConstr.of_constr concl in match_with_equation sigma concl >>= fun with_eqn -> match with_eqn with | Some eq_data,_,_ -> @@ -5003,7 +4980,7 @@ let abstract_subproof id gk tac = else (Context.Named.add d s1,s2)) global_sign (Context.Named.empty, empty_named_context_val) in let id = next_global_ident_away id (pf_ids_of_hyps gl) in - let concl = it_mkNamedProd_or_LetIn (EConstr.of_constr (Proofview.Goal.concl gl)) sign in + let concl = it_mkNamedProd_or_LetIn (Proofview.Goal.concl gl) sign in let concl = try flush_and_check_evars !evdref concl with Uninstantiated_evar _ -> diff --git a/toplevel/auto_ind_decl.ml b/toplevel/auto_ind_decl.ml index f8ca8343c..5f33ae834 100644 --- a/toplevel/auto_ind_decl.ml +++ b/toplevel/auto_ind_decl.ml @@ -338,8 +338,6 @@ so from Ai we can find the the correct eq_Ai bl_ai or lb_ai *) (* used in the leib -> bool side*) let do_replace_lb mode lb_scheme_key aavoid narg p q = - let p = EConstr.of_constr p in - let q = EConstr.of_constr q in let avoid = Array.of_list aavoid in let do_arg v offset = try @@ -399,6 +397,8 @@ let do_replace_lb mode lb_scheme_key aavoid narg p q = (* used in the bool -> leib side *) let do_replace_bl mode bl_scheme_key (ind,u as indu) aavoid narg lft rgt = + let lft = EConstr.Unsafe.to_constr lft in + let rgt = EConstr.Unsafe.to_constr rgt in let avoid = Array.of_list aavoid in let do_arg v offset = try @@ -612,11 +612,12 @@ repeat ( apply andb_prop in z;let z1:= fresh "Z" in destruct z as [z1 z]). Ci a1 ... an = Ci b1 ... bn replace bi with ai; auto || replace bi with ai by apply typeofbi_prod ; auto *) - Proofview.Goal.nf_enter { enter = begin fun gls -> - let gl = Proofview.Goal.concl gls in - match (kind_of_term gl) with + Proofview.Goal.nf_enter { enter = begin fun gl -> + let concl = Proofview.Goal.concl gl in + let sigma = Tacmach.New.project gl in + match EConstr.kind sigma concl with | App (c,ca) -> ( - match (kind_of_term c) with + match EConstr.kind sigma c with | Ind (indeq, u) -> if eq_gr (IndRef indeq) Coqlib.glob_eq then @@ -708,6 +709,7 @@ let compute_lb_goal ind lnamesparrec nparrec = ))), eff let compute_lb_tact mode lb_scheme_key ind lnamesparrec nparrec = + let open EConstr in let list_id = list_id lnamesparrec in let avoid = ref [] in let first_intros = @@ -745,10 +747,11 @@ let compute_lb_tact mode lb_scheme_key ind lnamesparrec nparrec = simplest_split ;Auto.default_auto ] ); Proofview.Goal.nf_enter { enter = begin fun gls -> - let gl = Proofview.Goal.concl gls in + let concl = Proofview.Goal.concl gls in + let sigma = Tacmach.New.project gl in (* assume the goal to be eq (eq_type ...) = true *) - match (kind_of_term gl) with - | App(c,ca) -> (match (kind_of_term ca.(1)) with + match EConstr.kind sigma concl with + | App(c,ca) -> (match (EConstr.kind sigma ca.(1)) with | App(c',ca') -> let n = Array.length ca' in do_replace_lb mode lb_scheme_key -- cgit v1.2.3 From 0cdb7e42f64674e246d4e24e3c725e23ceeec6bd Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Mon, 21 Nov 2016 12:13:05 +0100 Subject: Reductionops now return EConstrs. --- engine/termops.ml | 3 + engine/termops.mli | 1 + interp/notation.ml | 2 +- kernel/term.ml | 2 - kernel/term.mli | 1 - ltac/extratactics.ml4 | 17 +++-- ltac/rewrite.ml | 19 +---- ltac/tacinterp.ml | 1 + plugins/btauto/refl_btauto.ml | 1 + plugins/decl_mode/decl_proof_instr.ml | 4 +- plugins/extraction/extraction.ml | 22 ++++-- plugins/firstorder/instances.ml | 2 +- plugins/firstorder/unify.ml | 4 +- plugins/funind/functional_principles_proofs.ml | 28 ++++--- plugins/funind/functional_principles_types.ml | 1 + plugins/funind/invfun.ml | 12 +-- plugins/funind/recdef.ml | 3 +- plugins/omega/coq_omega.ml | 10 +-- plugins/romega/const_omega.ml | 2 +- plugins/setoid_ring/newring.ml | 4 +- pretyping/cases.ml | 11 +-- pretyping/classops.ml | 2 - pretyping/coercion.ml | 22 +++--- pretyping/evardefine.ml | 8 +- pretyping/evarsolve.ml | 10 +-- pretyping/indrec.ml | 6 +- pretyping/inductiveops.ml | 10 +-- pretyping/nativenorm.ml | 2 +- pretyping/nativenorm.mli | 2 +- pretyping/pretyping.ml | 4 +- pretyping/recordops.ml | 8 +- pretyping/reductionops.ml | 79 ++++++++++--------- pretyping/reductionops.mli | 101 ++++++++++++------------- pretyping/retyping.ml | 12 +-- pretyping/tacred.ml | 51 ++++++------- pretyping/typeclasses.ml | 1 + pretyping/typing.ml | 7 +- pretyping/unification.ml | 23 +++--- pretyping/vnorm.ml | 2 +- pretyping/vnorm.mli | 2 +- printing/prettyp.ml | 4 +- proofs/clenv.ml | 8 +- proofs/logic.ml | 12 +-- proofs/logic.mli | 2 +- proofs/redexpr.ml | 2 +- proofs/tacmach.ml | 8 +- proofs/tacmach.mli | 14 ++-- tactics/autorewrite.ml | 1 + tactics/class_tactics.ml | 1 - tactics/contradiction.ml | 3 +- tactics/eauto.ml | 2 +- tactics/eqschemes.ml | 4 +- tactics/equality.ml | 10 +-- tactics/hints.ml | 3 +- tactics/leminv.ml | 4 +- tactics/tactics.ml | 74 +++++++++--------- tactics/tactics.mli | 2 +- toplevel/command.ml | 4 +- toplevel/himsg.ml | 5 +- toplevel/obligations.ml | 4 +- toplevel/record.ml | 1 + toplevel/vernacentries.ml | 2 +- 62 files changed, 328 insertions(+), 344 deletions(-) diff --git a/engine/termops.ml b/engine/termops.ml index abaa409bd..879d77de2 100644 --- a/engine/termops.ml +++ b/engine/termops.ml @@ -692,6 +692,9 @@ 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) diff --git a/engine/termops.mli b/engine/termops.mli index 426b5f34d..1699d600e 100644 --- a/engine/termops.mli +++ b/engine/termops.mli @@ -130,6 +130,7 @@ val occur_term : Evd.evar_map -> EConstr.t -> EConstr.t -> bool (** Synonymous o (* Substitution of metavariables *) type meta_value_map = (metavariable * constr) list val subst_meta : meta_value_map -> constr -> constr +val isMetaOf : Evd.evar_map -> metavariable -> EConstr.constr -> bool (** Type assignment for metavariables *) type meta_type_map = (metavariable * types) list diff --git a/interp/notation.ml b/interp/notation.ml index d264a1904..29e5a3bfd 100644 --- a/interp/notation.ml +++ b/interp/notation.ml @@ -615,7 +615,7 @@ let find_scope_class_opt = function (* Special scopes associated to arguments of a global reference *) let rec compute_arguments_classes t = - match kind_of_term (Reductionops.whd_betaiotazeta Evd.empty (EConstr.of_constr t)) with + match kind_of_term (EConstr.Unsafe.to_constr (Reductionops.whd_betaiotazeta Evd.empty (EConstr.of_constr t))) with | Prod (_,t,u) -> let cl = try Some (compute_scope_class t) with Not_found -> None in cl :: compute_arguments_classes u diff --git a/kernel/term.ml b/kernel/term.ml index 62c161be4..3881cd12d 100644 --- a/kernel/term.ml +++ b/kernel/term.ml @@ -179,8 +179,6 @@ let destMeta c = match kind_of_term c with | _ -> raise DestKO let isMeta c = match kind_of_term c with Meta _ -> true | _ -> false -let isMetaOf mv c = - match kind_of_term c with Meta mv' -> Int.equal mv mv' | _ -> false (* Destructs a variable *) let destVar c = match kind_of_term c with diff --git a/kernel/term.mli b/kernel/term.mli index a8d9dfbff..e393adb81 100644 --- a/kernel/term.mli +++ b/kernel/term.mli @@ -98,7 +98,6 @@ val isVarId : Id.t -> constr -> bool val isInd : constr -> bool val isEvar : constr -> bool val isMeta : constr -> bool -val isMetaOf : metavariable -> constr -> bool val isEvar_or_Meta : constr -> bool val isSort : constr -> bool val isCast : constr -> bool diff --git a/ltac/extratactics.ml4 b/ltac/extratactics.ml4 index 188d041c1..138400991 100644 --- a/ltac/extratactics.ml4 +++ b/ltac/extratactics.ml4 @@ -290,6 +290,7 @@ END (* Hint Resolve *) open Term +open EConstr open Vars open Coqlib @@ -298,23 +299,26 @@ let project_hint pri l2r r = let env = Global.env() in let sigma = Evd.from_env env in let sigma, c = Evd.fresh_global env sigma gr in - let t = Retyping.get_type_of env sigma (EConstr.of_constr c) in + let c = EConstr.of_constr c in + let t = Retyping.get_type_of env sigma c in + let t = EConstr.of_constr t in let t = - Tacred.reduce_to_quantified_ref env sigma (Lazy.force coq_iff_ref) (EConstr.of_constr t) in - let t = EConstr.Unsafe.to_constr t in - let sign,ccl = decompose_prod_assum t in - let (a,b) = match snd (decompose_app ccl) with + Tacred.reduce_to_quantified_ref env sigma (Lazy.force coq_iff_ref) t in + let sign,ccl = decompose_prod_assum sigma t in + let (a,b) = match snd (decompose_app sigma ccl) with | [a;b] -> (a,b) | _ -> assert false in let p = if l2r then build_coq_iff_left_proj () else build_coq_iff_right_proj () in - let c = Reductionops.whd_beta Evd.empty (EConstr.of_constr (mkApp (c, Context.Rel.to_extended_vect 0 sign))) in + let p = EConstr.of_constr p in + let c = Reductionops.whd_beta sigma (mkApp (c, Array.map EConstr.of_constr (Context.Rel.to_extended_vect 0 sign))) in let c = it_mkLambda_or_LetIn (mkApp (p,[|mkArrow a (lift 1 b);mkArrow b (lift 1 a);c|])) sign in let id = Nameops.add_suffix (Nametab.basename_of_global gr) ("_proj_" ^ (if l2r then "l2r" else "r2l")) in let ctx = Evd.universe_context_set sigma in + let c = EConstr.to_constr sigma c in let c = Declare.declare_definition ~internal:Declare.InternalTacticRequest id (c,ctx) in (pri,false,true,Hints.PathAny, Hints.IsGlobRef (Globnames.ConstRef c)) @@ -751,7 +755,6 @@ let mkCaseEq a : unit Proofview.tactic = let env = Proofview.Goal.env gl in (** FIXME: this looks really wrong. Does anybody really use this tactic? *) let Sigma (c, _, _) = (Tacred.pattern_occs [Locus.OnlyOccurrences [1], a]).Reductionops.e_redfun env (Sigma.Unsafe.of_evar_map Evd.empty) concl in - let c = EConstr.of_constr c in change_concl c end }; simplest_case a] diff --git a/ltac/rewrite.ml b/ltac/rewrite.ml index 4c350b093..c92ddf990 100644 --- a/ltac/rewrite.ml +++ b/ltac/rewrite.ml @@ -247,14 +247,11 @@ end) = struct in let rec aux env evars ty l = let t = Reductionops.whd_all env (goalevars evars) ty in - let t = EConstr.of_constr t in match EConstr.kind (goalevars evars) t, l with | Prod (na, ty, b), obj :: cstrs -> let b = Reductionops.nf_betaiota (goalevars evars) b in - let b = EConstr.of_constr b in if noccurn (goalevars evars) 1 b (* non-dependent product *) then let ty = Reductionops.nf_betaiota (goalevars evars) ty in - let ty = EConstr.of_constr ty in let (evars, b', arg, cstrs) = aux env evars (subst1 mkProp b) cstrs in let evars, relty = mk_relty evars env ty obj in let evars, newarg = app_poly env evars respectful [| ty ; b' ; relty ; arg |] in @@ -264,7 +261,6 @@ end) = struct aux (Environ.push_rel (local_assum (na, ty)) env) evars b cstrs in let ty = Reductionops.nf_betaiota (goalevars evars) ty in - let ty = EConstr.of_constr ty in let pred = mkLambda (na, ty, b) in let liftarg = mkLambda (na, ty, arg) in let evars, arg' = app_poly env evars forall_relation [| ty ; pred ; liftarg |] in @@ -275,7 +271,6 @@ end) = struct (match finalcstr with | None | Some (_, None) -> let t = Reductionops.nf_betaiota (fst evars) ty in - let t = EConstr.of_constr t in let evars, rel = mk_relty evars env t None in evars, t, rel, [t, Some rel] | Some (t, Some rel) -> evars, t, rel, [t, Some rel]) @@ -355,7 +350,7 @@ end) = struct if Int.equal n 0 then start evars env prod else let sigma = goalevars evars in - match EConstr.kind sigma (EConstr.of_constr (Reductionops.whd_all env sigma prod)) with + match EConstr.kind sigma (Reductionops.whd_all env sigma prod) with | Prod (na, ty, b) -> if noccurn sigma 1 b then let b' = lift (-1) b in @@ -375,7 +370,6 @@ end) = struct with Not_found -> let sigma = goalevars evars in let ty = Reductionops.whd_all env sigma ty in - let ty = EConstr.of_constr ty in find env (mkApp (c, [| arg |])) (prod_applist sigma ty [arg]) args in find env c ty args @@ -513,7 +507,6 @@ let error_no_relation () = error "Cannot find a relation to rewrite." let rec decompose_app_rel env evd t = (** Head normalize for compatibility with the old meta mechanism *) let t = Reductionops.whd_betaiota evd t in - let t = EConstr.of_constr t in match EConstr.kind evd t with | App (f, [||]) -> assert false | App (f, [|arg|]) -> @@ -1011,7 +1004,7 @@ let unfold_match env sigma sk app = | App (f', args) when eq_constant (fst (destConst sigma f')) sk -> let v = Environ.constant_value_in (Global.env ()) (sk,Univ.Instance.empty)(*FIXME*) in let v = EConstr.of_constr v in - EConstr.of_constr (Reductionops.whd_beta sigma (mkApp (v, args))) + Reductionops.whd_beta sigma (mkApp (v, args)) | _ -> app let is_rew_cast = function RewCast _ -> true | _ -> false @@ -1106,7 +1099,7 @@ let subterm all flags (s : 'a pure_strategy) : 'a pure_strategy = | x -> x in let res = - { rew_car = EConstr.of_constr (Reductionops.hnf_prod_appvect env (goalevars evars) r.rew_car args); + { rew_car = Reductionops.hnf_prod_appvect env (goalevars evars) r.rew_car args; rew_from = mkApp(r.rew_from, args); rew_to = mkApp(r.rew_to, args); rew_prf = prf; rew_evars = r.rew_evars } in @@ -1455,7 +1448,6 @@ module Strategies = let rfn, ckind = Redexpr.reduction_of_red_expr env r in let sigma = Sigma.Unsafe.of_evar_map (goalevars evars) in let Sigma (t', sigma, _) = rfn.Reductionops.e_redfun env sigma t in - let t' = EConstr.of_constr t' in let evars' = Sigma.to_evar_map sigma in if Termops.eq_constr evars' t' t then state, Identity @@ -1475,7 +1467,6 @@ module Strategies = with e when CErrors.noncritical e -> error "fold: the term is not unfoldable !" in - let unfolded = EConstr.of_constr unfolded in try let sigma = Unification.w_unify env sigma CONV ~flags:(Unification.elim_flags ()) unfolded t in let c' = nf_evar sigma c in @@ -1568,7 +1559,6 @@ let cl_rewrite_clause_aux ?(abs=None) strat env avoid sigma concl is_hyp : resul | RewCast c -> None | RewPrf (rel, p) -> let p = nf_zeta env evars' (nf_evar evars' p) in - let p = EConstr.of_constr p in let term = match abs with | None -> p @@ -1951,8 +1941,7 @@ let declare_projection n instance_id r = | _ -> typ in aux init in - let ctx,ccl = Reductionops.splay_prod_n env sigma (3 * n) typ in - let ccl = EConstr.of_constr ccl + let ctx,ccl = Reductionops.splay_prod_n env sigma (3 * n) typ in it_mkProd_or_LetIn ccl ctx in let typ = it_mkProd_or_LetIn typ ctx in diff --git a/ltac/tacinterp.ml b/ltac/tacinterp.ml index 572fa32b7..6c59abe66 100644 --- a/ltac/tacinterp.ml +++ b/ltac/tacinterp.ml @@ -785,6 +785,7 @@ let interp_may_eval f ist env sigma = function let (redfun, _) = Redexpr.reduction_of_red_expr env redexp in let sigma = Sigma.Unsafe.of_evar_map sigma in let Sigma (c, sigma, _) = redfun.Reductionops.e_redfun env sigma (EConstr.of_constr c_interp) in + let c = EConstr.Unsafe.to_constr c in (Sigma.to_evar_map sigma, c) | ConstrContext ((loc,s),c) -> (try diff --git a/plugins/btauto/refl_btauto.ml b/plugins/btauto/refl_btauto.ml index 93bd88fe4..a0b04ce3b 100644 --- a/plugins/btauto/refl_btauto.ml +++ b/plugins/btauto/refl_btauto.ml @@ -182,6 +182,7 @@ module Btauto = struct let var = EConstr.of_constr var in (* Compute an assignment that dissatisfies the goal *) let _, var = Tacmach.pf_reduction_of_red_expr gl (Genredexpr.CbvVm None) var in + let var = EConstr.Unsafe.to_constr var in let rec to_list l = match decomp_term (Tacmach.project gl) l with | Term.App (c, _) when c === (Lazy.force CoqList._nil) -> [] diff --git a/plugins/decl_mode/decl_proof_instr.ml b/plugins/decl_mode/decl_proof_instr.ml index c3254010a..7123ebcaf 100644 --- a/plugins/decl_mode/decl_proof_instr.ml +++ b/plugins/decl_mode/decl_proof_instr.ml @@ -1103,7 +1103,7 @@ let thesis_for obj typ per_info env= ((Printer.pr_constr_env env Evd.empty obj) ++ spc () ++ str "cannot give an induction hypothesis (wrong parameters).") in let hd2 = (applist ((lift (List.length rc) per_info.per_pred),args@[obj])) in - compose_prod rc (Reductionops.whd_beta Evd.empty (EConstr.of_constr hd2)) + compose_prod rc (EConstr.Unsafe.to_constr (Reductionops.whd_beta Evd.empty (EConstr.of_constr hd2))) let rec build_product_dep pat_info per_info args body gls = match args with @@ -1233,7 +1233,7 @@ let hrec_for fix_id per_info gls obj_id = try List.for_all2 Term.eq_constr params per_info.per_params with Invalid_argument _ -> false end; let hd2 = applist (mkVar fix_id,args@[obj]) in - EConstr.of_constr (compose_lam rc (Reductionops.whd_beta gls.sigma (EConstr.of_constr hd2))) + EConstr.of_constr (compose_lam rc (EConstr.Unsafe.to_constr (Reductionops.whd_beta gls.sigma (EConstr.of_constr hd2)))) let warn_missing_case = CWarnings.create ~name:"declmode-missing-case" ~category:"declmode" diff --git a/plugins/extraction/extraction.ml b/plugins/extraction/extraction.ml index 460157130..61547f96d 100644 --- a/plugins/extraction/extraction.ml +++ b/plugins/extraction/extraction.ml @@ -70,11 +70,17 @@ type scheme = TypeScheme | Default type flag = info * scheme +let whd_all env t = + EConstr.Unsafe.to_constr (whd_all env none (EConstr.of_constr t)) + +let whd_betaiotazeta t = + EConstr.Unsafe.to_constr (whd_betaiotazeta none (EConstr.of_constr t)) + (*s [flag_of_type] transforms a type [t] into a [flag]. Really important function. *) let rec flag_of_type env t : flag = - let t = whd_all env none (EConstr.of_constr t) in + let t = whd_all env t in match kind_of_term t with | Prod (x,t,c) -> flag_of_type (push_rel (LocalAssum (x,t)) env) c | Sort s when Sorts.is_prop s -> (Logic,TypeScheme) @@ -105,14 +111,14 @@ let push_rel_assum (n, t) env = (*s [type_sign] gernerates a signature aimed at treating a type application. *) let rec type_sign env c = - match kind_of_term (whd_all env none (EConstr.of_constr c)) with + match kind_of_term (whd_all env c) with | Prod (n,t,d) -> (if is_info_scheme env t then Keep else Kill Kprop) :: (type_sign (push_rel_assum (n,t) env) d) | _ -> [] let rec type_scheme_nb_args env c = - match kind_of_term (whd_all env none (EConstr.of_constr c)) with + match kind_of_term (whd_all env c) with | Prod (n,t,d) -> let n = type_scheme_nb_args (push_rel_assum (n,t) env) d in if is_info_scheme env t then n+1 else n @@ -138,7 +144,7 @@ let make_typvar n vl = next_ident_away id' vl let rec type_sign_vl env c = - match kind_of_term (whd_all env none (EConstr.of_constr c)) with + match kind_of_term (whd_all env c) with | Prod (n,t,d) -> let s,vl = type_sign_vl (push_rel_assum (n,t) env) d in if not (is_info_scheme env t) then Kill Kprop::s, vl @@ -146,7 +152,7 @@ let rec type_sign_vl env c = | _ -> [],[] let rec nb_default_params env c = - match kind_of_term (whd_all env none (EConstr.of_constr c)) with + match kind_of_term (whd_all env c) with | Prod (n,t,d) -> let n = nb_default_params (push_rel_assum (n,t) env) d in if is_default env t then n+1 else n @@ -217,7 +223,7 @@ let parse_ind_args si args relmax = let rec extract_type env db j c args = - match kind_of_term (whd_betaiotazeta none (EConstr.of_constr c)) with + match kind_of_term (whd_betaiotazeta c) with | App (d, args') -> (* We just accumulate the arguments. *) extract_type env db j d (Array.to_list args' @ args) @@ -319,7 +325,7 @@ and extract_type_app env db (r,s) args = and extract_type_scheme env db c p = if Int.equal p 0 then extract_type env db 0 c [] else - let c = whd_betaiotazeta none (EConstr.of_constr c) in + let c = whd_betaiotazeta c in match kind_of_term c with | Lambda (n,t,d) -> extract_type_scheme (push_rel_assum (n,t) env) db d (p-1) @@ -492,7 +498,7 @@ and extract_really_ind env kn mib = *) and extract_type_cons env db dbmap c i = - match kind_of_term (whd_all env none (EConstr.of_constr c)) with + match kind_of_term (whd_all env c) with | Prod (n,t,d) -> let env' = push_rel_assum (n,t) env in let db' = (try Int.Map.find i dbmap with Not_found -> 0) :: db in diff --git a/plugins/firstorder/instances.ml b/plugins/firstorder/instances.ml index a320b47aa..24d4346d9 100644 --- a/plugins/firstorder/instances.ml +++ b/plugins/firstorder/instances.ml @@ -107,7 +107,7 @@ let mk_open_instance id idc gl m t= let typ=pf_unsafe_type_of gl (EConstr.of_constr idc) in (* since we know we will get a product, reduction is not too expensive *) - let (nam,_,_)=destProd (whd_all env evmap (EConstr.of_constr typ)) in + let (nam,_,_)=destProd (EConstr.Unsafe.to_constr (whd_all env evmap (EConstr.of_constr typ))) in match nam with Name id -> id | Anonymous -> dummy_bvid in diff --git a/plugins/firstorder/unify.ml b/plugins/firstorder/unify.ml index 205cb282d..5520c7e35 100644 --- a/plugins/firstorder/unify.ml +++ b/plugins/firstorder/unify.ml @@ -42,8 +42,8 @@ let unif t1 t2= Queue.add (t1,t2) bige; try while true do let t1,t2=Queue.take bige in - let nt1=head_reduce (whd_betaiotazeta Evd.empty (EConstr.of_constr t1)) - and nt2=head_reduce (whd_betaiotazeta Evd.empty (EConstr.of_constr t2)) in + let nt1=head_reduce (EConstr.Unsafe.to_constr (whd_betaiotazeta evd (EConstr.of_constr t1))) + and nt2=head_reduce (EConstr.Unsafe.to_constr (whd_betaiotazeta evd (EConstr.of_constr t2))) in match (kind_of_term nt1),(kind_of_term nt2) with Meta i,Meta j-> if not (Int.equal i j) then diff --git a/plugins/funind/functional_principles_proofs.ml b/plugins/funind/functional_principles_proofs.ml index 2e3992be9..188368082 100644 --- a/plugins/funind/functional_principles_proofs.ml +++ b/plugins/funind/functional_principles_proofs.ml @@ -321,6 +321,7 @@ let change_eq env sigma hyp_id (context:Context.Rel.t) x t end_of_type = in let new_type_of_hyp = Reductionops.nf_betaiota Evd.empty (EConstr.of_constr new_type_of_hyp) in + let new_type_of_hyp = EConstr.Unsafe.to_constr new_type_of_hyp in let new_ctxt,new_end_of_type = decompose_prod_n_assum ctxt_size new_type_of_hyp in @@ -619,6 +620,7 @@ let treat_new_case ptes_infos nb_prod continue_tac term dyn_infos = ) in let new_body = pf_nf_betaiota g' (EConstr.of_constr (mkApp(fun_body,[| new_term_value |]))) in + let new_body = EConstr.Unsafe.to_constr new_body in let new_infos = {dyn_infos with info = new_body; @@ -752,6 +754,7 @@ let build_proof pf_nf_betaiota g' (EConstr.of_constr (mkApp(dyn_infos.info,[|mkVar id|]))) in + let new_term = EConstr.Unsafe.to_constr new_term in let new_infos = {dyn_infos with info = new_term} in let do_prove new_hyps = build_proof do_finalize @@ -796,6 +799,7 @@ let build_proof | Lambda _ -> let new_term = Reductionops.nf_beta Evd.empty (EConstr.of_constr dyn_infos.info) in + let new_term = EConstr.Unsafe.to_constr new_term in build_proof do_finalize {dyn_infos with info = new_term} g | LetIn _ -> @@ -1097,11 +1101,11 @@ let prove_princ_for_struct (evd:Evd.evar_map ref) interactive_proof fun_num fnam let get_body const = match Global.body_of_constant const with | Some body -> - Tacred.cbv_norm_flags + EConstr.Unsafe.to_constr (Tacred.cbv_norm_flags (CClosure.RedFlags.mkflags [CClosure.RedFlags.fZETA]) (Global.env ()) (Evd.empty) - (EConstr.of_constr body) + (EConstr.of_constr body)) | None -> error ( "Cannot define a principle over an axiom ") in let fbody = get_body fnames.(fun_num) in @@ -1152,9 +1156,9 @@ let prove_princ_for_struct (evd:Evd.evar_map ref) interactive_proof fun_num fnam let bodies_with_all_params = Array.map (fun body -> - Reductionops.nf_betaiota Evd.empty + EConstr.Unsafe.to_constr (Reductionops.nf_betaiota Evd.empty (EConstr.of_constr (applist(substl (List.rev (Array.to_list all_funs_with_full_params)) body, - List.rev_map var_of_decl princ_params))) + List.rev_map var_of_decl princ_params)))) ) bodies in @@ -1190,12 +1194,12 @@ let prove_princ_for_struct (evd:Evd.evar_map ref) interactive_proof fun_num fnam let body_with_param,num = let body = get_body fnames.(i) in let body_with_full_params = - Reductionops.nf_betaiota Evd.empty (EConstr.of_constr ( - applist(body,List.rev_map var_of_decl full_params))) + EConstr.Unsafe.to_constr (Reductionops.nf_betaiota Evd.empty (EConstr.of_constr ( + applist(body,List.rev_map var_of_decl full_params)))) in match kind_of_term body_with_full_params with | Fix((_,num),(_,_,bs)) -> - Reductionops.nf_betaiota Evd.empty + EConstr.Unsafe.to_constr (Reductionops.nf_betaiota Evd.empty (EConstr.of_constr ( (applist (substl @@ -1203,7 +1207,7 @@ let prove_princ_for_struct (evd:Evd.evar_map ref) interactive_proof fun_num fnam (Array.to_list all_funs_with_full_params)) bs.(num), List.rev_map var_of_decl princ_params)) - )),num + ))),num | _ -> error "Not a mutual block" in let info = @@ -1279,8 +1283,8 @@ let prove_princ_for_struct (evd:Evd.evar_map ref) interactive_proof fun_num fnam nb_rec_hyps = -100; rec_hyps = []; info = - Reductionops.nf_betaiota Evd.empty - (EConstr.of_constr (applist(fix_body,List.rev_map mkVar args_id))); + EConstr.Unsafe.to_constr (Reductionops.nf_betaiota Evd.empty + (EConstr.of_constr (applist(fix_body,List.rev_map mkVar args_id)))); eq_hyps = [] } in @@ -1339,11 +1343,11 @@ let prove_princ_for_struct (evd:Evd.evar_map ref) interactive_proof fun_num fnam nb_rec_hyps = -100; rec_hyps = []; info = - Reductionops.nf_betaiota Evd.empty + EConstr.Unsafe.to_constr (Reductionops.nf_betaiota Evd.empty (EConstr.of_constr (applist(fbody_with_full_params, (List.rev_map var_of_decl princ_params)@ (List.rev_map mkVar args_id) - ))); + )))); eq_hyps = [] } in diff --git a/plugins/funind/functional_principles_types.ml b/plugins/funind/functional_principles_types.ml index 4d88f9f91..b4eb77870 100644 --- a/plugins/funind/functional_principles_types.ml +++ b/plugins/funind/functional_principles_types.ml @@ -409,6 +409,7 @@ let get_funs_constant mp dp = (Evd.from_env (Global.env ())) (EConstr.of_constr body) in + let body = EConstr.Unsafe.to_constr body in body | None -> error ( "Cannot define a principle over an axiom ") in diff --git a/plugins/funind/invfun.ml b/plugins/funind/invfun.ml index d29d4694f..c02b64c1f 100644 --- a/plugins/funind/invfun.ml +++ b/plugins/funind/invfun.ml @@ -252,7 +252,6 @@ let prove_fun_correct evd functional_induction funs_constr graphs_constr schemes (* and the principle to use in this lemma in $\zeta$ normal form *) let f_principle,princ_type = schemes.(i) in let princ_type = nf_zeta (EConstr.of_constr princ_type) in - let princ_type = EConstr.of_constr princ_type in let princ_infos = Tactics.compute_elim_sig evd princ_type in (* The number of args of the function is then easily computable *) let nb_fun_args = nb_prod (project g) (EConstr.of_constr (pf_concl g)) - 2 in @@ -429,7 +428,7 @@ let prove_fun_correct evd functional_induction funs_constr graphs_constr schemes List.rev (fst (List.fold_left2 (fun (bindings,avoid) decl p -> let id = Namegen.next_ident_away (Nameops.out_name (RelDecl.get_name decl)) avoid in - (nf_zeta (EConstr.of_constr p))::bindings,id::avoid) + (EConstr.Unsafe.to_constr (nf_zeta (EConstr.of_constr p)))::bindings,id::avoid) ([],avoid) princ_infos.predicates (lemmas))) @@ -662,7 +661,7 @@ let prove_fun_complete funcs graphs schemes lemmas_types_infos i : tactic = (* We get the constant and the principle corresponding to this lemma *) let f = funcs.(i) in let graph_principle = nf_zeta (EConstr.of_constr schemes.(i)) in - let princ_type = pf_unsafe_type_of g (EConstr.of_constr graph_principle) in + let princ_type = pf_unsafe_type_of g graph_principle in let princ_type = EConstr.of_constr princ_type in let princ_infos = Tactics.compute_elim_sig (project g) princ_type in (* Then we get the number of argument of the function @@ -754,14 +753,15 @@ let prove_fun_complete funcs graphs schemes lemmas_types_infos i : tactic = g in let params_names = fst (List.chop princ_infos.nparams args_names) in + let open EConstr in let params = List.map mkVar params_names in tclTHENSEQ [ tclMAP (fun id -> Proofview.V82.of_tactic (Simple.intro id)) (args_names@[res;hres]); observe_tac "h_generalize" - (Proofview.V82.of_tactic (generalize [EConstr.of_constr (mkApp(applist(graph_principle,params),Array.map (fun c -> applist(c,params)) lemmas))])); + (Proofview.V82.of_tactic (generalize [mkApp(applist(graph_principle,params),Array.map (fun c -> applist(c,params)) lemmas)])); Proofview.V82.of_tactic (Simple.intro graph_principle_id); observe_tac "" (tclTHEN_i - (observe_tac "elim" (Proofview.V82.of_tactic (elim false None (EConstr.mkVar hres,NoBindings) (Some (EConstr.mkVar graph_principle_id,NoBindings))))) + (observe_tac "elim" (Proofview.V82.of_tactic (elim false None (mkVar hres,NoBindings) (Some (mkVar graph_principle_id,NoBindings))))) (fun i g -> observe_tac "prove_branche" (prove_branche i) g )) ] g @@ -796,6 +796,7 @@ let derive_correctness make_scheme functional_induction (funs: pconstant list) ( let type_of_lemma = Termops.it_mkProd_or_LetIn type_of_lemma_concl type_of_lemma_ctxt in let _ = Typing.e_type_of (Global.env ()) evd (EConstr.of_constr type_of_lemma) in let type_of_lemma = nf_zeta (EConstr.of_constr type_of_lemma) in + let type_of_lemma = EConstr.Unsafe.to_constr type_of_lemma in observe (str "type_of_lemma := " ++ Printer.pr_lconstr_env (Global.env ()) !evd type_of_lemma); type_of_lemma,type_info ) @@ -863,6 +864,7 @@ let derive_correctness make_scheme functional_induction (funs: pconstant list) ( Termops.it_mkProd_or_LetIn type_of_lemma_concl type_of_lemma_ctxt in let type_of_lemma = nf_zeta (EConstr.of_constr type_of_lemma) in + let type_of_lemma = EConstr.Unsafe.to_constr type_of_lemma in observe (str "type_of_lemma := " ++ Printer.pr_lconstr type_of_lemma); type_of_lemma,type_info ) diff --git a/plugins/funind/recdef.ml b/plugins/funind/recdef.ml index 5cee3cb20..c71174fef 100644 --- a/plugins/funind/recdef.ml +++ b/plugins/funind/recdef.ml @@ -695,7 +695,7 @@ let mkDestructEq : let changefun patvars = { run = fun sigma -> let redfun = pattern_occs [Locus.AllOccurrencesBut [1], EConstr.of_constr expr] in let Sigma (c, sigma, p) = redfun.Reductionops.e_redfun (pf_env g2) sigma (EConstr.of_constr (pf_concl g2)) in - Sigma (EConstr.of_constr c, sigma, p) + Sigma (c, sigma, p) } in Proofview.V82.of_tactic (change_in_concl None changefun) g2); Proofview.V82.of_tactic (simplest_case (EConstr.of_constr expr))]), to_revert @@ -1503,6 +1503,7 @@ let recursive_definition is_mes function_name rec_impls type_of_f r rec_arg_num let res_vars,eq' = decompose_prod equation_lemma_type in let env_eq' = Environ.push_rel_context (List.map (fun (x,y) -> LocalAssum (x,y)) res_vars) env in let eq' = nf_zeta env_eq' (EConstr.of_constr eq') in + let eq' = EConstr.Unsafe.to_constr eq' in let res = (* Pp.msgnl (str "res_var :=" ++ Printer.pr_lconstr_env (push_rel_context (List.map (function (x,t) -> (x,None,t)) res_vars) env) eq'); *) (* Pp.msgnl (str "rec_arg_num := " ++ str (string_of_int rec_arg_num)); *) diff --git a/plugins/omega/coq_omega.ml b/plugins/omega/coq_omega.ml index 72290e681..51790f4c9 100644 --- a/plugins/omega/coq_omega.ml +++ b/plugins/omega/coq_omega.ml @@ -568,7 +568,7 @@ let abstract_path typ path t = mkLambda (Name (Id.of_string "x"), typ, abstract), !term_occur let focused_simpl path gl = - let newc = context (fun i t -> pf_nf gl (EConstr.of_constr t)) (List.rev path) (pf_concl gl) in + let newc = context (fun i t -> EConstr.Unsafe.to_constr (pf_nf gl (EConstr.of_constr t))) (List.rev path) (pf_concl gl) in let newc = EConstr.of_constr newc in Proofview.V82.of_tactic (convert_concl_no_check newc DEFAULTcast) gl @@ -1386,7 +1386,7 @@ let destructure_omega gl tac_def (id,c) = else try match destructurate_prop c with | Kapp(Eq,[typ;t1;t2]) - when begin match destructurate_type (pf_nf gl (EConstr.of_constr typ)) with Kapp(Z,[]) -> true | _ -> false end -> + when begin match destructurate_type (EConstr.Unsafe.to_constr (pf_nf gl (EConstr.of_constr typ))) with Kapp(Z,[]) -> true | _ -> false end -> let t = mk_plus t1 (mk_inv t2) in normalize_equation id EQUA (Lazy.force coq_Zegal_left) 2 t t1 t2 tac_def @@ -1661,7 +1661,7 @@ let rec decidability gl t = | Kapp(Not,[t1]) -> mkApp (Lazy.force coq_dec_not, [| t1; decidability gl t1 |]) | Kapp(Eq,[typ;t1;t2]) -> - begin match destructurate_type (pf_nf gl (EConstr.of_constr typ)) with + begin match destructurate_type (EConstr.Unsafe.to_constr (pf_nf gl (EConstr.of_constr typ))) with | Kapp(Z,[]) -> mkApp (Lazy.force coq_dec_eq, [| t1;t2 |]) | Kapp(Nat,[]) -> mkApp (Lazy.force coq_dec_eq_nat, [| t1;t2 |]) | _ -> raise Undecidable @@ -1791,7 +1791,7 @@ let destructure_hyps = with Not_found -> loop lit) | Kapp(Eq,[typ;t1;t2]) -> if !old_style_flag then begin - match destructurate_type (pf_nf (EConstr.of_constr typ)) with + match destructurate_type (EConstr.Unsafe.to_constr (pf_nf (EConstr.of_constr typ))) with | Kapp(Nat,_) -> Tacticals.New.tclTHENLIST [ (simplest_elim @@ -1808,7 +1808,7 @@ let destructure_hyps = ] | _ -> loop lit end else begin - match destructurate_type (pf_nf (EConstr.of_constr typ)) with + match destructurate_type (EConstr.Unsafe.to_constr (pf_nf (EConstr.of_constr typ))) with | Kapp(Nat,_) -> (Tacticals.New.tclTHEN (convert_hyp_no_check (NamedDecl.set_type (mkApp (Lazy.force coq_neq, [| t1;t2|])) diff --git a/plugins/romega/const_omega.ml b/plugins/romega/const_omega.ml index f2d91bad3..5c68078d7 100644 --- a/plugins/romega/const_omega.ml +++ b/plugins/romega/const_omega.ml @@ -353,7 +353,7 @@ let parse_term t = let parse_rel gl t = try match destructurate t with | Kapp("eq",[typ;t1;t2]) - when destructurate (Tacmach.pf_nf gl (EConstr.of_constr typ)) = Kapp("Z",[]) -> Req (t1,t2) + when destructurate (EConstr.Unsafe.to_constr (Tacmach.pf_nf gl (EConstr.of_constr typ))) = Kapp("Z",[]) -> Req (t1,t2) | Kapp("Zne",[t1;t2]) -> Rne (t1,t2) | Kapp("Z.le",[t1;t2]) -> Rle (t1,t2) | Kapp("Z.lt",[t1;t2]) -> Rlt (t1,t2) diff --git a/plugins/setoid_ring/newring.ml b/plugins/setoid_ring/newring.ml index 4492b854b..b720b2e0a 100644 --- a/plugins/setoid_ring/newring.ml +++ b/plugins/setoid_ring/newring.ml @@ -83,8 +83,8 @@ let lookup_map map = let protect_red map env sigma c = let c = EConstr.Unsafe.to_constr c in - kl (create_clos_infos all env) - (mk_clos_but (lookup_map map c) (Esubst.subs_id 0) c);; + EConstr.of_constr (kl (create_clos_infos all env) + (mk_clos_but (lookup_map map c) (Esubst.subs_id 0) c));; let protect_tac map = Tactics.reduct_option (protect_red map,DEFAULTcast) None diff --git a/pretyping/cases.ml b/pretyping/cases.ml index 76ced2b1d..c0141f011 100644 --- a/pretyping/cases.ml +++ b/pretyping/cases.ml @@ -1057,7 +1057,6 @@ let specialize_predicate newtomatchs (names,depna) arsign cs tms ccl = (* Note: applying the substitution in tms is not important (is it sure?) *) let ccl'' = whd_betaiota Evd.empty (subst_predicate (realargsi, copti) ccl' tms) in - let ccl'' = EConstr.of_constr ccl'' in (* We adjust ccl st: gamma, x'1..x'n, x1..xn, tms |- ccl'' *) let ccl''' = liftn_predicate n (n+1) ccl'' tms in (* We finally get gamma,x'1..x'n,x |- [X1;x1:I(X1)]..[Xn;xn:I(Xn)]pred'''*) @@ -1065,8 +1064,8 @@ let specialize_predicate newtomatchs (names,depna) arsign cs tms ccl = let find_predicate loc env evdref p current (IndType (indf,realargs)) dep tms = let pred = abstract_predicate env !evdref indf current realargs dep tms p in - (pred, EConstr.of_constr (whd_betaiota !evdref - (applist (pred, realargs@[current])))) + (pred, whd_betaiota !evdref + (applist (pred, realargs@[current]))) (* Take into account that a type has been discovered to be inductive, leading to more dependencies in the predicate if the type has indices *) @@ -1221,7 +1220,7 @@ let rec generalize_problem names pb = function | LocalDef (Anonymous,_,_) -> pb', deps | _ -> (* for better rendering *) - let d = RelDecl.map_type (fun c -> whd_betaiota !(pb.evdref) (EConstr.of_constr c)) d in + let d = RelDecl.map_type (fun c -> EConstr.Unsafe.to_constr (whd_betaiota !(pb.evdref) (EConstr.of_constr c))) d in let tomatch = lift_tomatch_stack 1 pb'.tomatch in let tomatch = relocate_index_tomatch !(pb.evdref) (i+1) 1 tomatch in { pb' with @@ -1400,7 +1399,6 @@ and match_current pb (initial,tomatch) = pred current indt (names,dep) tomatch in let ci = make_case_info pb.env (fst mind) pb.casestyle in let pred = nf_betaiota !(pb.evdref) pred in - let pred = EConstr.of_constr pred in let case = make_case_or_project pb.env !(pb.evdref) indf ci pred current brvals in @@ -1638,7 +1636,6 @@ let rec list_assoc_in_triple x = function let abstract_tycon loc env evdref subst tycon extenv t = let t = nf_betaiota !evdref t in (* it helps in some cases to remove K-redex*) - let t = EConstr.of_constr t in let src = match EConstr.kind !evdref t with | Evar (evk,_) -> (loc,Evar_kinds.SubEvar evk) | _ -> (loc,Evar_kinds.CasesType true) in @@ -1734,7 +1731,7 @@ let build_inversion_problem loc env sigma tms t = let id = next_name_away (named_hd env (EConstr.Unsafe.to_constr t) Anonymous) avoid in PatVar (Loc.ghost,Name id), ((id,t)::subst, id::avoid) in let rec reveal_pattern t (subst,avoid as acc) = - match EConstr.kind sigma (EConstr.of_constr (whd_all env sigma t)) with + match EConstr.kind sigma (whd_all env sigma t) with | Construct (cstr,u) -> PatCstr (Loc.ghost,cstr,[],Anonymous), acc | App (f,v) when isConstruct sigma f -> let cstr,u = destConstruct sigma f in diff --git a/pretyping/classops.ml b/pretyping/classops.ml index 23d20dad3..e4331aade 100644 --- a/pretyping/classops.ml +++ b/pretyping/classops.ml @@ -234,7 +234,6 @@ let class_of env sigma t = (t, n1, i, u, args) with Not_found -> let t = Tacred.hnf_constr env sigma t in - let t = EConstr.of_constr t in let (cl, u, args) = find_class_type sigma t in let (i, { cl_param = n1 } ) = class_info cl in (t, n1, i, u, args) @@ -279,7 +278,6 @@ let apply_on_class_of env sigma t cont = with Not_found -> (* Is it worth to be more incremental on the delta steps? *) let t = Tacred.hnf_constr env sigma t in - let t = EConstr.of_constr t in let (cl, u, args) = find_class_type sigma t in let (i, { cl_param = n1 } ) = class_info cl in if not (Int.equal (List.length args) n1) then raise Not_found; diff --git a/pretyping/coercion.ml b/pretyping/coercion.ml index 48f7be4bb..7e8559630 100644 --- a/pretyping/coercion.ml +++ b/pretyping/coercion.ml @@ -64,7 +64,7 @@ let apply_coercion_args env evd check isproj argl funj = { uj_val = applist (j_val funj,argl); uj_type = typ } | h::restl -> (* On devrait pouvoir s'arranger pour qu'on n'ait pas a faire hnf_constr *) - match EConstr.kind !evdref (EConstr.of_constr (whd_all env !evdref typ)) with + match EConstr.kind !evdref (whd_all env !evdref typ) with | Prod (_,c1,c2) -> if check && not (e_cumul env evdref (EConstr.of_constr (Retyping.get_type_of env !evdref h)) c1) then raise NoCoercion; @@ -96,7 +96,7 @@ let make_existential loc ?(opaque = not (get_proofs_transparency ())) env evdre Evarutil.e_new_evar env evdref ~src c let app_opt env evdref f t = - EConstr.of_constr (whd_betaiota !evdref (app_opt f t)) + whd_betaiota !evdref (app_opt f t) let pair_of_array a = (a.(0), a.(1)) @@ -134,11 +134,10 @@ let local_assum (na, t) = let mu env evdref t = let rec aux v = let v' = hnf env !evdref v in - let v' = EConstr.of_constr v' in match disc_subset !evdref v' with | Some (u, p) -> let f, ct = aux u in - let p = EConstr.of_constr (hnf_nodelta env !evdref p) in + let p = hnf_nodelta env !evdref p in (Some (fun x -> app_opt env evdref f (papp evdref sig_proj1 [| u; p; x |])), @@ -152,8 +151,6 @@ and coerce loc env evdref (x : EConstr.constr) (y : EConstr.constr) let open Context.Rel.Declaration in let rec coerce_unify env x y = let x = hnf env !evdref x and y = hnf env !evdref y in - let x = EConstr.of_constr x in - let y = EConstr.of_constr y in try evdref := the_conv_x_leq env x y !evdref; None @@ -162,7 +159,7 @@ and coerce loc env evdref (x : EConstr.constr) (y : EConstr.constr) let subco () = subset_coerce env evdref x y in let dest_prod c = match Reductionops.splay_prod_n env (!evdref) 1 c with - | [LocalAssum (na,t) | LocalDef (na,_,t)], c -> (na, EConstr.of_constr t), EConstr.of_constr c + | [LocalAssum (na,t) | LocalDef (na,_,t)], c -> (na, EConstr.of_constr t), c | _ -> raise NoSubtacCoercion in let coerce_application typ typ' c c' l l' = @@ -344,7 +341,7 @@ let app_coercion env evdref coercion v = | None -> v | Some f -> let v' = Typing.e_solve_evars env evdref (f v) in - EConstr.of_constr (whd_betaiota !evdref (EConstr.of_constr v')) + whd_betaiota !evdref (EConstr.of_constr v') let coerce_itf loc env evd v t c1 = let evdref = ref evd in @@ -381,7 +378,6 @@ let apply_coercion env sigma p hj typ_cl = (* Try to coerce to a funclass; raise NoCoercion if not possible *) let inh_app_fun_core env evd j = let t = whd_all env evd j.uj_type in - let t = EConstr.of_constr t in match EConstr.kind evd t with | Prod (_,_,_) -> (evd,j) | Evar ev -> @@ -413,7 +409,7 @@ let inh_app_fun resolve_tc env evd j = with NoCoercion -> (evd, j) let type_judgment env sigma j = - match EConstr.kind sigma (EConstr.of_constr (whd_all env sigma j.uj_type)) with + match EConstr.kind sigma (whd_all env sigma j.uj_type) with | Sort s -> {utj_val = j.uj_val; utj_type = s } | _ -> error_not_a_type env sigma j @@ -429,7 +425,7 @@ let inh_tosort_force loc env evd j = let inh_coerce_to_sort loc env evd j = let typ = whd_all env evd j.uj_type in - match EConstr.kind evd (EConstr.of_constr typ) with + match EConstr.kind evd typ with | Sort s -> (evd,{ utj_val = j.uj_val; utj_type = s }) | Evar ev -> let (evd',s) = Evardefine.define_evar_as_sort env evd ev in @@ -480,8 +476,8 @@ let rec inh_conv_coerce_to_fail loc env evd rigidonly v t c1 = try inh_coerce_to_fail env evd rigidonly v t c1 with NoCoercion -> match - EConstr.kind evd (EConstr.of_constr (whd_all env evd t)), - EConstr.kind evd (EConstr.of_constr (whd_all env evd c1)) + EConstr.kind evd (whd_all env evd t), + EConstr.kind evd (whd_all env evd c1) with | Prod (name,t1,t2), Prod (_,u1,u2) -> (* Conversion did not work, we may succeed with a coercion. *) diff --git a/pretyping/evardefine.ml b/pretyping/evardefine.ml index 3babc48a7..d4b46c046 100644 --- a/pretyping/evardefine.ml +++ b/pretyping/evardefine.ml @@ -39,7 +39,7 @@ let env_nf_evar sigma env = let env_nf_betaiotaevar sigma env = process_rel_context (fun d e -> - push_rel (RelDecl.map_constr (fun c -> Reductionops.nf_betaiota sigma (EConstr.of_constr c)) d) e) env + push_rel (RelDecl.map_constr (fun c -> EConstr.Unsafe.to_constr (Reductionops.nf_betaiota sigma (EConstr.of_constr c))) d) e) env (****************************************) (* Operations on value/type constraints *) @@ -85,7 +85,6 @@ let define_pure_evar_as_product evd evk = let evenv = evar_env evi in let id = next_ident_away idx (ids_of_named_context (evar_context evi)) in let concl = Reductionops.whd_all evenv evd (EConstr.of_constr evi.evar_concl) in - let concl = EConstr.of_constr concl in let s = destSort evd concl in let evd1,(dom,u1) = let evd = Sigma.Unsafe.of_evar_map evd in @@ -138,7 +137,7 @@ let define_pure_evar_as_lambda env evd evk = let open Context.Named.Declaration in let evi = Evd.find_undefined evd evk in let evenv = evar_env evi in - let typ = EConstr.of_constr (Reductionops.whd_all evenv evd (EConstr.of_constr (evar_concl evi))) in + let typ = Reductionops.whd_all evenv evd (EConstr.of_constr (evar_concl evi)) in let evd1,(na,dom,rng) = match EConstr.kind evd typ with | Prod (na,dom,rng) -> (evd,(na,dom,rng)) | Evar ev' -> let evd,typ = define_evar_as_product evd ev' in evd,destProd evd typ @@ -177,7 +176,6 @@ let define_evar_as_sort env evd (ev,args) = let evi = Evd.find_undefined evd ev in let s = Type u in let concl = Reductionops.whd_all (evar_env evi) evd (EConstr.of_constr evi.evar_concl) in - let concl = EConstr.of_constr concl in let sort = destSort evd concl in let evd' = Evd.define ev (Constr.mkSort s) evd in Evd.set_leq_sort env evd' (Type (Univ.super u)) sort, s @@ -190,7 +188,7 @@ let define_evar_as_sort env evd (ev,args) = let split_tycon loc env evd tycon = let rec real_split evd c = let t = Reductionops.whd_all env evd c in - match EConstr.kind evd (EConstr.of_constr t) with + match EConstr.kind evd t with | Prod (na,dom,rng) -> evd, (na, dom, rng) | Evar ev (* ev is undefined because of whd_all *) -> let (evd',prod) = define_evar_as_product evd ev in diff --git a/pretyping/evarsolve.ml b/pretyping/evarsolve.ml index 65b6d287d..27436fdd8 100644 --- a/pretyping/evarsolve.ml +++ b/pretyping/evarsolve.ml @@ -149,7 +149,7 @@ let recheck_applications conv_algo env evdref t = let argsty = Array.map (fun x -> aux env x; EConstr.of_constr (Retyping.get_type_of env !evdref x)) args in let rec aux i ty = if i < Array.length argsty then - match EConstr.kind !evdref (EConstr.of_constr (whd_all env !evdref ty)) with + match EConstr.kind !evdref (whd_all env !evdref ty) with | Prod (na, dom, codom) -> (match conv_algo env !evdref Reduction.CUMUL argsty.(i) dom with | Success evd -> evdref := evd; @@ -814,7 +814,7 @@ let rec do_projection_effects define_fun env ty evd = function let evd = Evd.define evk (Constr.mkVar id) evd in (* TODO: simplify constraints involving evk *) let evd = do_projection_effects define_fun env ty evd p in - let ty = EConstr.of_constr (whd_all env evd (Lazy.force ty)) in + let ty = whd_all env evd (Lazy.force ty) in if not (isSort evd ty) then (* Don't try to instantiate if a sort because if evar_concl is an evar it may commit to a univ level which is not the right @@ -1494,7 +1494,7 @@ let rec invert_definition conv_algo choose env evd pbty (evk,argsv as ev) rhs = map_constr_with_full_binders !evdref (fun d (env,k) -> push_rel d env, k+1) imitate envk t in - let rhs = EConstr.of_constr (whd_beta evd rhs) (* heuristic *) in + let rhs = whd_beta evd rhs (* heuristic *) in let fast rhs = let filter_ctxt = evar_filtered_context evi in let names = ref Idset.empty in @@ -1576,7 +1576,7 @@ and evar_define conv_algo ?(choose=false) env evd pbty (evk,argsv as ev) rhs = raise e | OccurCheckIn (evd,rhs) -> (* last chance: rhs actually reduces to ev *) - let c = EConstr.of_constr (whd_all env evd rhs) in + let c = whd_all env evd rhs in match EConstr.kind evd c with | Evar (evk',argsv2) when Evar.equal evk evk' -> solve_refl (fun env sigma pb c c' -> is_fconv pb env sigma c c') @@ -1637,7 +1637,7 @@ let reconsider_conv_pbs conv_algo evd = (* Rq: uncomplete algorithm if pbty = CONV_X_LEQ ! *) let solve_simple_eqn conv_algo ?(choose=false) env evd (pbty,(evk1,args1 as ev1),t2) = try - let t2 = EConstr.of_constr (whd_betaiota evd t2) in (* includes whd_evar *) + let t2 = whd_betaiota evd t2 in (* includes whd_evar *) let evd = evar_define conv_algo ~choose env evd pbty ev1 t2 in reconsider_conv_pbs conv_algo evd with diff --git a/pretyping/indrec.ml b/pretyping/indrec.ml index 4fa5ad06d..1adeb4db2 100644 --- a/pretyping/indrec.ml +++ b/pretyping/indrec.ml @@ -173,6 +173,7 @@ let type_rec_branch is_rec dep env sigma (vargs,depPvect,decP) tyi cs recargs = base | _ -> let t' = whd_all env sigma (EConstr.of_constr p) in + let t' = EConstr.Unsafe.to_constr t' in if Term.eq_constr p' t' then assert false else prec env i sign t' in @@ -247,6 +248,7 @@ let make_rec_branch_arg env sigma (nparrec,fvect,decF) f cstr recargs = applist(lift i fk,realargs@[arg]) | _ -> let t' = whd_all env sigma (EConstr.of_constr p) in + let t' = EConstr.Unsafe.to_constr t' in if Term.eq_constr t' p' then assert false else prec env i hyps t' in @@ -265,7 +267,7 @@ let make_rec_branch_arg env sigma (nparrec,fvect,decF) f cstr recargs = | None -> mkLambda_name env (n,t,process_constr (push_rel d env) (i+1) - (whd_beta Evd.empty (EConstr.of_constr (applist (lift 1 f, [(mkRel 1)])))) + (EConstr.Unsafe.to_constr (whd_beta Evd.empty (EConstr.of_constr (applist (lift 1 f, [(mkRel 1)]))))) (cprest,rest)) | Some(_,f_0) -> let nF = lift (i+1+decF) f_0 in @@ -273,7 +275,7 @@ let make_rec_branch_arg env sigma (nparrec,fvect,decF) f cstr recargs = let arg = process_pos env' nF (lift 1 t) in mkLambda_name env (n,t,process_constr env' (i+1) - (whd_beta Evd.empty (EConstr.of_constr (applist (lift 1 f, [(mkRel 1); arg])))) + (EConstr.Unsafe.to_constr (whd_beta Evd.empty (EConstr.of_constr (applist (lift 1 f, [(mkRel 1); arg]))))) (cprest,rest))) | (LocalDef (n,c,t) as d)::cprest, rest -> mkLetIn diff --git a/pretyping/inductiveops.ml b/pretyping/inductiveops.ml index 9c5a2e894..120adb9fe 100644 --- a/pretyping/inductiveops.ml +++ b/pretyping/inductiveops.ml @@ -456,7 +456,7 @@ let extract_mrectype sigma t = let find_mrectype_vect env sigma c = let open EConstr in - let (t, l) = Termops.decompose_app_vect sigma (EConstr.of_constr (whd_all env sigma c)) in + let (t, l) = Termops.decompose_app_vect sigma (whd_all env sigma c) in match EConstr.kind sigma t with | Ind ind -> (ind, l) | _ -> raise Not_found @@ -466,7 +466,7 @@ let find_mrectype env sigma c = let find_rectype env sigma c = let open EConstr in - let (t, l) = decompose_app sigma (EConstr.of_constr (whd_all env sigma c)) in + let (t, l) = decompose_app sigma (whd_all env sigma c) in match EConstr.kind sigma t with | Ind (ind,u as indu) -> let (mib,mip) = Inductive.lookup_mind_specif env ind in @@ -478,7 +478,7 @@ let find_rectype env sigma c = let find_inductive env sigma c = let open EConstr in - let (t, l) = decompose_app sigma (EConstr.of_constr (whd_all env sigma c)) in + let (t, l) = decompose_app sigma (whd_all env sigma c) in match EConstr.kind sigma t with | Ind ind when (fst (Inductive.lookup_mind_specif env (fst ind))).mind_finite <> Decl_kinds.CoFinite -> @@ -488,7 +488,7 @@ let find_inductive env sigma c = let find_coinductive env sigma c = let open EConstr in - let (t, l) = decompose_app sigma (EConstr.of_constr (whd_all env sigma c)) in + let (t, l) = decompose_app sigma (whd_all env sigma c) in match EConstr.kind sigma t with | Ind ind when (fst (Inductive.lookup_mind_specif env (fst ind))).mind_finite == Decl_kinds.CoFinite -> @@ -503,7 +503,7 @@ let find_coinductive env sigma c = let is_predicate_explicitly_dep env sigma pred arsign = let rec srec env pval arsign = - let pv' = EConstr.of_constr (whd_all env sigma pval) in + let pv' = whd_all env sigma pval in match EConstr.kind sigma pv', arsign with | Lambda (na,t,b), (LocalAssum _)::arsign -> srec (push_rel_assum (na, t) env) b arsign diff --git a/pretyping/nativenorm.ml b/pretyping/nativenorm.ml index cdaa4e9ee..0228f63cd 100644 --- a/pretyping/nativenorm.ml +++ b/pretyping/nativenorm.ml @@ -404,7 +404,7 @@ let native_norm env sigma c ty = let t2 = Sys.time () in let time_info = Format.sprintf "Reification done in %.5f@." (t2 -. t1) in if !Flags.debug then Feedback.msg_debug (Pp.str time_info); - res + EConstr.of_constr res | _ -> anomaly (Pp.str "Compilation failure") let native_conv_generic pb sigma t = diff --git a/pretyping/nativenorm.mli b/pretyping/nativenorm.mli index ba46138a4..c899340c8 100644 --- a/pretyping/nativenorm.mli +++ b/pretyping/nativenorm.mli @@ -12,7 +12,7 @@ open Evd (** This module implements normalization by evaluation to OCaml code *) -val native_norm : env -> evar_map -> constr -> types -> Constr.t +val native_norm : env -> evar_map -> constr -> types -> constr (** Conversion with inference of universe constraints *) val native_infer_conv : ?pb:conv_pb -> env -> evar_map -> constr -> constr -> diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml index f814028f9..7d2c96bb9 100644 --- a/pretyping/pretyping.ml +++ b/pretyping/pretyping.ml @@ -743,7 +743,6 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre let argloc = loc_of_glob_constr c in let resj = evd_comb1 (Coercion.inh_app_fun resolve_tc env.ExtraEnv.env) evdref resj in let resty = whd_all env.ExtraEnv.env !evdref resj.uj_type in - let resty = EConstr.of_constr resty in match EConstr.kind !evdref resty with | Prod (na,c1,c2) -> let tycon = Some c1 in @@ -917,7 +916,6 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre @[EConstr.of_constr (build_dependent_constructor cs)] in let lp = lift cs.cs_nargs p in let fty = hnf_lam_applist env.ExtraEnv.env !evdref lp inst in - let fty = EConstr.of_constr fty in let fj = pretype (mk_tycon fty) env_f evdref lvar d in let v = let ind,_ = dest_ind_family indf in @@ -1100,7 +1098,7 @@ and pretype_type k0 resolve_tc valcon (env : ExtraEnv.t) evdref lvar = function let sigma = !evdref in let t = Retyping.get_type_of env.ExtraEnv.env sigma v in let t = EConstr.of_constr t in - match EConstr.kind sigma (EConstr.of_constr (whd_all env.ExtraEnv.env sigma t)) with + match EConstr.kind sigma (whd_all env.ExtraEnv.env sigma t) with | Sort s -> s | Evar ev when is_Type (existential_type sigma ev) -> evd_comb1 (define_evar_as_sort env.ExtraEnv.env) evdref ev diff --git a/pretyping/recordops.ml b/pretyping/recordops.ml index 3230f92da..8362a2a26 100644 --- a/pretyping/recordops.ml +++ b/pretyping/recordops.ml @@ -324,15 +324,15 @@ let lookup_canonical_conversion (proj,pat) = assoc_pat pat (Refmap.find proj !object_table) let is_open_canonical_projection env sigma (c,args) = + let open EConstr in try - let c = EConstr.Unsafe.to_constr c in - let ref = global_of_constr c in + let (ref, _) = Termops.global_of_constr sigma c in let n = find_projection_nparams ref in (** Check if there is some canonical projection attached to this structure *) let _ = Refmap.find ref !object_table in try let arg = whd_all env sigma (Stack.nth args n) in - let hd = match kind_of_term arg with App (hd, _) -> hd | _ -> arg in - not (isConstruct hd) + let hd = match EConstr.kind sigma arg with App (hd, _) -> hd | _ -> arg in + not (isConstruct sigma hd) with Failure _ -> false with Not_found -> false diff --git a/pretyping/reductionops.ml b/pretyping/reductionops.ml index 6ec3cd985..45e7abcb7 100644 --- a/pretyping/reductionops.ml +++ b/pretyping/reductionops.ml @@ -588,10 +588,10 @@ end (** The type of (machine) states (= lambda-bar-calculus' cuts) *) type state = constr * constr Stack.t -type contextual_reduction_function = env -> evar_map -> constr -> Constr.constr +type contextual_reduction_function = env -> evar_map -> constr -> constr type reduction_function = contextual_reduction_function -type local_reduction_function = evar_map -> constr -> Constr.constr -type e_reduction_function = { e_redfun : 'r. env -> 'r Sigma.t -> constr -> (Constr.constr, 'r) Sigma.sigma } +type local_reduction_function = evar_map -> constr -> constr +type e_reduction_function = { e_redfun : 'r. env -> 'r Sigma.t -> constr -> (constr, 'r) Sigma.sigma } type contextual_stack_reduction_function = env -> evar_map -> constr -> constr * constr list @@ -629,19 +629,18 @@ let safe_meta_value sigma ev = let strong whdfun env sigma t = let rec strongrec env t = - let t = EConstr.of_constr (whdfun env sigma t) in - map_constr_with_full_binders sigma push_rel strongrec env t in - EConstr.Unsafe.to_constr (strongrec env t) + map_constr_with_full_binders sigma push_rel strongrec env (whdfun env sigma t) in + strongrec env t let local_strong whdfun sigma = - let rec strongrec t = EConstr.map sigma strongrec (EConstr.of_constr (whdfun sigma t)) in - fun c -> EConstr.Unsafe.to_constr (strongrec c) + let rec strongrec t = EConstr.map sigma strongrec (whdfun sigma t) in + strongrec let rec strong_prodspine redfun sigma c = - let x = EConstr.of_constr (redfun sigma c) in + let x = redfun sigma c in match EConstr.kind sigma x with - | Prod (na,a,b) -> EConstr.Unsafe.to_constr (mkProd (na,a,EConstr.of_constr (strong_prodspine redfun sigma b))) - | _ -> EConstr.Unsafe.to_constr x + | Prod (na,a,b) -> mkProd (na,a,strong_prodspine redfun sigma b) + | _ -> x (*************************************) (*** Reduction using bindingss ***) @@ -1140,7 +1139,7 @@ let iterate_whd_gen refold flags env sigma s = in aux s let red_of_state_red f sigma x = - EConstr.Unsafe.to_constr (Stack.zip sigma (f sigma (x,Stack.empty))) + Stack.zip sigma (f sigma (x,Stack.empty)) (* 0. No Reduction Functions *) @@ -1217,9 +1216,9 @@ let nf_evar = Evarutil.nf_evar let clos_norm_flags flgs env sigma t = try let evars ev = safe_evar_value sigma ev in - CClosure.norm_val + EConstr.of_constr (CClosure.norm_val (CClosure.create_clos_infos ~evars flgs env) - (CClosure.inject (EConstr.Unsafe.to_constr t)) + (CClosure.inject (EConstr.Unsafe.to_constr t))) with e when is_anomaly e -> error "Tried to normalize ill-typed term" let nf_beta = clos_norm_flags CClosure.beta (Global.env ()) @@ -1309,6 +1308,7 @@ let sigma_univ_state = let infer_conv_gen conv_fun ?(catch_incon=true) ?(pb=Reduction.CUMUL) ?(ts=full_transparent_state) env sigma x y = + (** FIXME *) let x = EConstr.Unsafe.to_constr x in let y = EConstr.Unsafe.to_constr y in try @@ -1352,8 +1352,8 @@ let vm_infer_conv ?(pb=Reduction.CUMUL) env t1 t2 = (********************************************************************) let whd_meta sigma c = match EConstr.kind sigma c with - | Meta p -> (try meta_value sigma p with Not_found -> EConstr.Unsafe.to_constr c) - | _ -> EConstr.Unsafe.to_constr c + | Meta p -> (try EConstr.of_constr (meta_value sigma p) with Not_found -> c) + | _ -> c let default_plain_instance_ident = Id.of_string "H" @@ -1431,26 +1431,26 @@ let instance sigma s c = * error message. *) let hnf_prod_app env sigma t n = - match EConstr.kind sigma (EConstr.of_constr (whd_all env sigma t)) with - | Prod (_,_,b) -> EConstr.Unsafe.to_constr (subst1 n b) + match EConstr.kind sigma (whd_all env sigma t) with + | Prod (_,_,b) -> subst1 n b | _ -> anomaly ~label:"hnf_prod_app" (Pp.str "Need a product") let hnf_prod_appvect env sigma t nl = - Array.fold_left (fun acc t -> hnf_prod_app env sigma (EConstr.of_constr acc) t) (EConstr.Unsafe.to_constr t) nl + Array.fold_left (fun acc t -> hnf_prod_app env sigma acc t) t nl let hnf_prod_applist env sigma t nl = - List.fold_left (fun acc t -> hnf_prod_app env sigma (EConstr.of_constr acc) t) (EConstr.Unsafe.to_constr t) nl + List.fold_left (fun acc t -> hnf_prod_app env sigma acc t) t nl let hnf_lam_app env sigma t n = - match EConstr.kind sigma (EConstr.of_constr (whd_all env sigma t)) with - | Lambda (_,_,b) -> EConstr.Unsafe.to_constr (subst1 n b) + match EConstr.kind sigma (whd_all env sigma t) with + | Lambda (_,_,b) -> subst1 n b | _ -> anomaly ~label:"hnf_lam_app" (Pp.str "Need an abstraction") let hnf_lam_appvect env sigma t nl = - Array.fold_left (fun acc t -> hnf_lam_app env sigma (EConstr.of_constr acc) t) (EConstr.Unsafe.to_constr t) nl + Array.fold_left (fun acc t -> hnf_lam_app env sigma acc t) t nl let hnf_lam_applist env sigma t nl = - List.fold_left (fun acc t -> hnf_lam_app env sigma (EConstr.of_constr acc) t) (EConstr.Unsafe.to_constr t) nl + List.fold_left (fun acc t -> hnf_lam_app env sigma acc t) t nl let bind_assum (na, t) = (na, t) @@ -1458,7 +1458,6 @@ let bind_assum (na, t) = let splay_prod env sigma = let rec decrec env m c = let t = whd_all env sigma c in - let t = EConstr.of_constr t in match EConstr.kind sigma t with | Prod (n,a,c0) -> decrec (push_rel (local_assum (n,a)) env) @@ -1470,7 +1469,6 @@ let splay_prod env sigma = let splay_lam env sigma = let rec decrec env m c = let t = whd_all env sigma c in - let t = EConstr.of_constr t in match EConstr.kind sigma t with | Lambda (n,a,c0) -> decrec (push_rel (local_assum (n,a)) env) @@ -1482,7 +1480,7 @@ let splay_lam env sigma = let splay_prod_assum env sigma = let rec prodec_rec env l c = let t = whd_allnolet env sigma c in - match EConstr.kind sigma (EConstr.of_constr t) with + match EConstr.kind sigma t with | Prod (x,t,c) -> prodec_rec (push_rel (local_assum (x,t)) env) (Context.Rel.add (local_assum (x,t)) l) c @@ -1491,9 +1489,9 @@ let splay_prod_assum env sigma = (Context.Rel.add (local_def (x,b,t)) l) c | Cast (c,_,_) -> prodec_rec env l c | _ -> - let t' = whd_all env sigma (EConstr.of_constr t) in - if Term.eq_constr t t' then l,t - else prodec_rec env l (EConstr.of_constr t') + let t' = whd_all env sigma t in + if EConstr.eq_constr sigma t t' then l,t + else prodec_rec env l t' in prodec_rec env Context.Rel.empty @@ -1506,8 +1504,8 @@ let splay_arity env sigma c = let sort_of_arity env sigma c = snd (splay_arity env sigma c) let splay_prod_n env sigma n = - let rec decrec env m ln c = if Int.equal m 0 then (ln, EConstr.Unsafe.to_constr c) else - match EConstr.kind sigma (EConstr.of_constr (whd_all env sigma c)) with + let rec decrec env m ln c = if Int.equal m 0 then (ln,c) else + match EConstr.kind sigma (whd_all env sigma c) with | Prod (n,a,c0) -> decrec (push_rel (local_assum (n,a)) env) (m-1) (Context.Rel.add (local_assum (n,a)) ln) c0 @@ -1516,8 +1514,8 @@ let splay_prod_n env sigma n = decrec env n Context.Rel.empty let splay_lam_n env sigma n = - let rec decrec env m ln c = if Int.equal m 0 then (ln, EConstr.Unsafe.to_constr c) else - match EConstr.kind sigma (EConstr.of_constr (whd_all env sigma c)) with + let rec decrec env m ln c = if Int.equal m 0 then (ln,c) else + match EConstr.kind sigma (whd_all env sigma c) with | Lambda (n,a,c0) -> decrec (push_rel (local_assum (n,a)) env) (m-1) (Context.Rel.add (local_assum (n,a)) ln) c0 @@ -1526,7 +1524,7 @@ let splay_lam_n env sigma n = decrec env n Context.Rel.empty let is_sort env sigma t = - match EConstr.kind sigma (EConstr.of_constr (whd_all env sigma t)) with + match EConstr.kind sigma (whd_all env sigma t) with | Sort s -> true | _ -> false @@ -1559,7 +1557,7 @@ let whd_betaiota_deltazeta_for_iota_state ts env sigma csts s = let find_conclusion env sigma = let rec decrec env c = let t = whd_all env sigma c in - match EConstr.kind sigma (EConstr.of_constr t) with + match EConstr.kind sigma t with | Prod (x,t,c0) -> decrec (push_rel (local_assum (x,t)) env) c0 | Lambda (x,t,c0) -> decrec (push_rel (local_assum (x,t)) env) c0 | t -> t @@ -1579,7 +1577,7 @@ let meta_value evd mv = match meta_opt_fvalue evd mv with | Some (b,_) -> let metas = Metamap.bind valrec b.freemetas in - EConstr.of_constr (instance evd metas (EConstr.of_constr b.rebus)) + instance evd metas (EConstr.of_constr b.rebus) | None -> mkMeta mv in valrec mv @@ -1589,7 +1587,7 @@ let meta_instance sigma b = if Metaset.is_empty fm then b.rebus else let c_sigma = Metamap.bind (fun mv -> meta_value sigma mv) fm in - EConstr.of_constr (instance sigma c_sigma b.rebus) + instance sigma c_sigma b.rebus let nf_meta sigma c = let c = EConstr.Unsafe.to_constr c in @@ -1609,7 +1607,6 @@ let meta_reducible_instance evd b = let metas = Metaset.fold fold fm Metamap.empty in let rec irec u = let u = whd_betaiota Evd.empty u (** FIXME *) in - let u = EConstr.of_constr u in match EConstr.kind evd u with | Case (ci,p,c,bl) when EConstr.isMeta evd (strip_outer_cast evd c) -> let m = destMeta evd (strip_outer_cast evd c) in @@ -1674,7 +1671,7 @@ let head_unfold_under_prod ts env sigma c = match EConstr.kind sigma h with | Const cst -> beta_app sigma (unfold cst, l) | _ -> c in - EConstr.Unsafe.to_constr (aux c) + aux c let betazetaevar_applist sigma n c l = let rec stacklam n env t stack = @@ -1684,4 +1681,4 @@ let betazetaevar_applist sigma n c l = | LetIn(_,b,_,c), _ -> stacklam (n-1) (substl env b::env) c stack | Evar _, _ -> applist (substl env t, stack) | _ -> anomaly (Pp.str "Not enough lambda/let's") in - EConstr.Unsafe.to_constr (stacklam n [] c l) + stacklam n [] c l diff --git a/pretyping/reductionops.mli b/pretyping/reductionops.mli index 3b3242537..add1d186b 100644 --- a/pretyping/reductionops.mli +++ b/pretyping/reductionops.mli @@ -8,6 +8,7 @@ open Names open Term +open EConstr open Univ open Evd open Environ @@ -38,7 +39,6 @@ val set_refolding_in_reduction : bool -> unit cst applied to params must convertible to term of the state applied to args *) module Cst_stack : sig - open EConstr type t val empty : t val add_param : constr -> t -> t @@ -52,7 +52,6 @@ end module Stack : sig - open EConstr type 'a app_node val pr_app_node : ('a -> Pp.std_ppcmds) -> 'a app_node -> Pp.std_ppcmds @@ -109,19 +108,19 @@ end (************************************************************************) -type state = EConstr.t * EConstr.t Stack.t +type state = constr * constr Stack.t -type contextual_reduction_function = env -> evar_map -> EConstr.t -> constr +type contextual_reduction_function = env -> evar_map -> constr -> constr type reduction_function = contextual_reduction_function -type local_reduction_function = evar_map -> EConstr.t -> constr +type local_reduction_function = evar_map -> constr -> constr -type e_reduction_function = { e_redfun : 'r. env -> 'r Sigma.t -> EConstr.t -> (constr, 'r) Sigma.sigma } +type e_reduction_function = { e_redfun : 'r. env -> 'r Sigma.t -> constr -> (constr, 'r) Sigma.sigma } type contextual_stack_reduction_function = - env -> evar_map -> EConstr.t -> EConstr.t * EConstr.t list + env -> evar_map -> constr -> constr * constr list type stack_reduction_function = contextual_stack_reduction_function type local_stack_reduction_function = - evar_map -> EConstr.t -> EConstr.t * EConstr.t list + evar_map -> constr -> constr * constr list type contextual_state_reduction_function = env -> evar_map -> state -> state @@ -139,13 +138,13 @@ val strong_prodspine : local_reduction_function -> local_reduction_function val stack_reduction_of_reduction : 'a reduction_function -> 'a state_reduction_function i*) -val stacklam : (state -> 'a) -> EConstr.t list -> evar_map -> EConstr.t -> EConstr.t Stack.t -> 'a +val stacklam : (state -> 'a) -> constr list -> evar_map -> constr -> constr Stack.t -> 'a val whd_state_gen : ?csts:Cst_stack.t -> refold:bool -> tactic_mode:bool -> CClosure.RedFlags.reds -> Environ.env -> Evd.evar_map -> state -> state * Cst_stack.t val iterate_whd_gen : bool -> CClosure.RedFlags.reds -> - Environ.env -> Evd.evar_map -> EConstr.constr -> EConstr.constr + Environ.env -> Evd.evar_map -> constr -> constr (** {6 Generic Optimized Reduction Function using Closures } *) @@ -156,11 +155,11 @@ val nf_beta : local_reduction_function val nf_betaiota : local_reduction_function val nf_betaiotazeta : local_reduction_function val nf_all : reduction_function -val nf_evar : evar_map -> constr -> constr +val nf_evar : evar_map -> Constr.constr -> Constr.constr (** Lazy strategy, weak head reduction *) -val whd_evar : evar_map -> constr -> constr +val whd_evar : evar_map -> Constr.constr -> Constr.constr val whd_nored : local_reduction_function val whd_beta : local_reduction_function val whd_betaiota : local_reduction_function @@ -198,45 +197,45 @@ val whd_zeta_stack : local_stack_reduction_function val whd_zeta_state : local_state_reduction_function val whd_zeta : local_reduction_function -val shrink_eta : EConstr.constr -> EConstr.constr +val shrink_eta : constr -> constr (** Various reduction functions *) -val safe_evar_value : evar_map -> existential -> constr option +val safe_evar_value : evar_map -> Constr.existential -> Constr.constr option -val beta_applist : evar_map -> EConstr.t * EConstr.t list -> EConstr.constr +val beta_applist : evar_map -> constr * constr list -> constr -val hnf_prod_app : env -> evar_map -> EConstr.t -> EConstr.t -> constr -val hnf_prod_appvect : env -> evar_map -> EConstr.t -> EConstr.t array -> constr -val hnf_prod_applist : env -> evar_map -> EConstr.t -> EConstr.t list -> constr -val hnf_lam_app : env -> evar_map -> EConstr.t -> EConstr.t -> constr -val hnf_lam_appvect : env -> evar_map -> EConstr.t -> EConstr.t array -> constr -val hnf_lam_applist : env -> evar_map -> EConstr.t -> EConstr.t list -> constr +val hnf_prod_app : env -> evar_map -> constr -> constr -> constr +val hnf_prod_appvect : env -> evar_map -> constr -> constr array -> constr +val hnf_prod_applist : env -> evar_map -> constr -> constr list -> constr +val hnf_lam_app : env -> evar_map -> constr -> constr -> constr +val hnf_lam_appvect : env -> evar_map -> constr -> constr array -> constr +val hnf_lam_applist : env -> evar_map -> constr -> constr list -> constr -val splay_prod : env -> evar_map -> EConstr.t -> (Name.t * EConstr.constr) list * EConstr.constr -val splay_lam : env -> evar_map -> EConstr.t -> (Name.t * EConstr.constr) list * EConstr.constr -val splay_arity : env -> evar_map -> EConstr.t -> (Name.t * EConstr.constr) list * sorts -val sort_of_arity : env -> evar_map -> EConstr.t -> sorts -val splay_prod_n : env -> evar_map -> int -> EConstr.t -> Context.Rel.t * constr -val splay_lam_n : env -> evar_map -> int -> EConstr.t -> Context.Rel.t * constr +val splay_prod : env -> evar_map -> constr -> (Name.t * constr) list * constr +val splay_lam : env -> evar_map -> constr -> (Name.t * constr) list * constr +val splay_arity : env -> evar_map -> constr -> (Name.t * constr) list * sorts +val sort_of_arity : env -> evar_map -> constr -> sorts +val splay_prod_n : env -> evar_map -> int -> constr -> Context.Rel.t * constr +val splay_lam_n : env -> evar_map -> int -> constr -> Context.Rel.t * constr val splay_prod_assum : - env -> evar_map -> EConstr.t -> Context.Rel.t * constr + env -> evar_map -> constr -> Context.Rel.t * constr type 'a miota_args = { - mP : EConstr.t; (** the result type *) - mconstr : EConstr.t; (** the constructor *) + mP : constr; (** the result type *) + mconstr : constr; (** the constructor *) mci : case_info; (** special info to re-build pattern *) mcargs : 'a list; (** the constructor's arguments *) mlf : 'a array } (** the branch code vector *) -val reducible_mind_case : evar_map -> EConstr.t -> bool -val reduce_mind_case : evar_map -> EConstr.t miota_args -> EConstr.t +val reducible_mind_case : evar_map -> constr -> bool +val reduce_mind_case : evar_map -> constr miota_args -> constr -val find_conclusion : env -> evar_map -> EConstr.t -> (EConstr.t,EConstr.t) kind_of_term -val is_arity : env -> evar_map -> EConstr.t -> bool -val is_sort : env -> evar_map -> EConstr.types -> bool +val find_conclusion : env -> evar_map -> constr -> (constr,constr) kind_of_term +val is_arity : env -> evar_map -> constr -> bool +val is_sort : env -> evar_map -> types -> bool -val contract_fix : ?env:Environ.env -> evar_map -> ?reference:Constant.t -> (EConstr.t, EConstr.t) pfixpoint -> EConstr.t +val contract_fix : ?env:Environ.env -> evar_map -> ?reference:Constant.t -> fixpoint -> constr val fix_recarg : ('a, 'a) pfixpoint -> 'b Stack.t -> (int * 'b) option (** {6 Querying the kernel conversion oracle: opaque/transparent constants } *) @@ -249,14 +248,14 @@ type conversion_test = constraints -> constraints val pb_is_equal : conv_pb -> bool val pb_equal : conv_pb -> conv_pb -val is_conv : ?reds:transparent_state -> env -> evar_map -> EConstr.t -> EConstr.t -> bool -val is_conv_leq : ?reds:transparent_state -> env -> evar_map -> EConstr.t -> EConstr.t -> bool -val is_fconv : ?reds:transparent_state -> conv_pb -> env -> evar_map -> EConstr.t -> EConstr.t -> bool +val is_conv : ?reds:transparent_state -> env -> evar_map -> constr -> constr -> bool +val is_conv_leq : ?reds:transparent_state -> env -> evar_map -> constr -> constr -> bool +val is_fconv : ?reds:transparent_state -> conv_pb -> env -> evar_map -> constr -> constr -> bool (** [check_conv] Checks universe constraints only. pb defaults to CUMUL and ts to a full transparent state. *) -val check_conv : ?pb:conv_pb -> ?ts:transparent_state -> env -> evar_map -> EConstr.t -> EConstr.t -> bool +val check_conv : ?pb:conv_pb -> ?ts:transparent_state -> env -> evar_map -> constr -> constr -> bool (** [infer_conv] Adds necessary universe constraints to the evar map. pb defaults to CUMUL and ts to a full transparent state. @@ -264,29 +263,29 @@ val check_conv : ?pb:conv_pb -> ?ts:transparent_state -> env -> evar_map -> ECo otherwise returns false in that case. *) val infer_conv : ?catch_incon:bool -> ?pb:conv_pb -> ?ts:transparent_state -> - env -> evar_map -> EConstr.constr -> EConstr.constr -> evar_map * bool + env -> evar_map -> constr -> constr -> evar_map * bool (** Conversion with inference of universe constraints *) -val set_vm_infer_conv : (?pb:conv_pb -> env -> evar_map -> EConstr.constr -> EConstr.constr -> +val set_vm_infer_conv : (?pb:conv_pb -> env -> evar_map -> constr -> constr -> evar_map * bool) -> unit -val vm_infer_conv : ?pb:conv_pb -> env -> evar_map -> EConstr.constr -> EConstr.constr -> +val vm_infer_conv : ?pb:conv_pb -> env -> evar_map -> constr -> constr -> evar_map * bool (** [infer_conv_gen] behaves like [infer_conv] but is parametrized by a conversion function. Used to pretype vm and native casts. *) val infer_conv_gen : (conv_pb -> l2r:bool -> evar_map -> transparent_state -> - (constr, evar_map) Reduction.generic_conversion_function) -> + (Constr.constr, evar_map) Reduction.generic_conversion_function) -> ?catch_incon:bool -> ?pb:conv_pb -> ?ts:transparent_state -> env -> - evar_map -> EConstr.constr -> EConstr.constr -> evar_map * bool + evar_map -> constr -> constr -> evar_map * bool (** {6 Special-Purpose Reduction Functions } *) val whd_meta : local_reduction_function -val plain_instance : evar_map -> EConstr.t Metamap.t -> EConstr.t -> EConstr.t -val instance : evar_map -> EConstr.t Metamap.t -> EConstr.t -> constr +val plain_instance : evar_map -> constr Metamap.t -> constr -> constr +val instance : evar_map -> constr Metamap.t -> constr -> constr val head_unfold_under_prod : transparent_state -> reduction_function -val betazetaevar_applist : evar_map -> int -> EConstr.t -> EConstr.t list -> constr +val betazetaevar_applist : evar_map -> int -> constr -> constr list -> constr (** {6 Heuristic for Conversion with Evar } *) @@ -295,6 +294,6 @@ val whd_betaiota_deltazeta_for_iota_state : state * Cst_stack.t (** {6 Meta-related reduction functions } *) -val meta_instance : evar_map -> EConstr.constr freelisted -> EConstr.constr -val nf_meta : evar_map -> EConstr.constr -> EConstr.constr -val meta_reducible_instance : evar_map -> EConstr.constr freelisted -> EConstr.constr +val meta_instance : evar_map -> constr freelisted -> constr +val nf_meta : evar_map -> constr -> constr +val meta_reducible_instance : evar_map -> constr freelisted -> constr diff --git a/pretyping/retyping.ml b/pretyping/retyping.ml index 3142ea5cb..7db30bf23 100644 --- a/pretyping/retyping.ml +++ b/pretyping/retyping.ml @@ -74,7 +74,7 @@ let get_type_from_constraints env sigma t = let rec subst_type env sigma typ = function | [] -> typ | h::rest -> - match EConstr.kind sigma (EConstr.of_constr (whd_all env sigma typ)) with + match EConstr.kind sigma (whd_all env sigma typ) with | Prod (na,c1,c2) -> subst_type env sigma (subst1 h c2) rest | _ -> retype_error NonFunctionalConstruction @@ -83,7 +83,7 @@ let rec subst_type env sigma typ = function let sort_of_atomic_type env sigma ft args = let rec concl_of_arity env n ar args = - match EConstr.kind sigma (EConstr.of_constr (whd_all env sigma ar)), args with + match EConstr.kind sigma (whd_all env sigma ar), args with | Prod (na, t, b), h::l -> concl_of_arity (push_rel (local_def (na, lift n h, t)) env) (n + 1) b l | Sort s, [] -> s | _ -> retype_error NotASort @@ -94,7 +94,7 @@ let type_of_var env id = with Not_found -> retype_error (BadVariable id) let decomp_sort env sigma t = - match EConstr.kind sigma (EConstr.of_constr (whd_all env sigma t)) with + match EConstr.kind sigma (whd_all env sigma t) with | Sort s -> s | _ -> retype_error NotASort @@ -123,9 +123,9 @@ let retype ?(polyprop=true) sigma = with Not_found -> retype_error BadRecursiveType in let n = inductive_nrealdecls_env env (fst (fst (dest_ind_family indf))) in - let t = EConstr.of_constr (betazetaevar_applist sigma n p realargs) in - (match EConstr.kind sigma (EConstr.of_constr (whd_all env sigma (type_of env t))) with - | Prod _ -> EConstr.of_constr (whd_beta sigma (applist (t, [c]))) + let t = betazetaevar_applist sigma n p realargs in + (match EConstr.kind sigma (whd_all env sigma (type_of env t)) with + | Prod _ -> whd_beta sigma (applist (t, [c])) | _ -> t) | Lambda (name,c1,c2) -> mkProd (name, c1, type_of (push_rel (local_assum (name,c1)) env) c2) diff --git a/pretyping/tacred.ml b/pretyping/tacred.ml index 1d179c683..02524f896 100644 --- a/pretyping/tacred.ml +++ b/pretyping/tacred.ml @@ -425,7 +425,7 @@ let solve_arity_problem env sigma fxminargs c = let evm = ref sigma in let set_fix i = evm := Evd.define i (Constr.mkVar vfx) !evm in let rec check strict c = - let c' = EConstr.of_constr (whd_betaiotazeta sigma c) in + let c' = whd_betaiotazeta sigma c in let (h,rcargs) = decompose_app_vect sigma c' in match EConstr.kind sigma h with Evar(i,_) when Evar.Map.mem i fxminargs && not (Evd.is_defined !evm i) -> @@ -473,7 +473,7 @@ let reduce_fix whdfun sigma fix stack = | None -> NotReducible | Some (recargnum,recarg) -> let (recarg'hd,_ as recarg') = whdfun sigma recarg in - let stack' = List.assign stack recargnum (EConstr.applist recarg') in + let stack' = List.assign stack recargnum (applist recarg') in (match EConstr.kind sigma recarg'hd with | Construct _ -> Reduced (contract_fix sigma fix, stack') | _ -> NotReducible) @@ -483,7 +483,7 @@ let contract_fix_use_function env sigma f let nbodies = Array.length recindices in let make_Fi j = (mkFix((recindices,j),typedbodies), f j) in let lbodies = List.init nbodies make_Fi in - substl_checking_arity env (List.rev lbodies) sigma (EConstr.of_constr (nf_beta sigma bodies.(bodynum))) + substl_checking_arity env (List.rev lbodies) sigma (nf_beta sigma bodies.(bodynum)) let reduce_fix_use_function env sigma f whfun fix stack = match fix_recarg fix (Stack.append_app_list stack Stack.empty) with @@ -495,7 +495,7 @@ let reduce_fix_use_function env sigma f whfun fix stack = (recarg, []) else whfun recarg in - let stack' = List.assign stack recargnum (EConstr.applist recarg') in + let stack' = List.assign stack recargnum (applist recarg') in (match EConstr.kind sigma recarg'hd with | Construct _ -> Reduced (contract_fix_use_function env sigma f fix,stack') @@ -507,7 +507,7 @@ let contract_cofix_use_function env sigma f let make_Fi j = (mkCoFix(j,typedbodies), f j) in let subbodies = List.init nbodies make_Fi in substl_checking_arity env (List.rev subbodies) - sigma (EConstr.of_constr (nf_beta sigma bodies.(bodynum))) + sigma (nf_beta sigma bodies.(bodynum)) let reduce_mind_case_use_function func env sigma mia = match EConstr.kind sigma mia.mconstr with @@ -689,7 +689,7 @@ let rec red_elim_const env sigma ref u largs = let whfun = whd_construct_stack env sigma in (match reduce_fix_use_function env sigma f whfun (destFix sigma d) lrest with | NotReducible -> raise Redelimination - | Reduced (c,rest) -> (EConstr.of_constr (nf_beta sigma c), rest), nocase) + | Reduced (c,rest) -> (nf_beta sigma c, rest), nocase) | EliminationMutualFix (min,refgoal,refinfos) when nargs >= min -> let rec descend (ref,u) args = let c = reference_value env sigma ref u in @@ -704,14 +704,14 @@ let rec red_elim_const env sigma ref u largs = let whfun = whd_construct_stack env sigma in (match reduce_fix_use_function env sigma f whfun (destFix sigma d) lrest with | NotReducible -> raise Redelimination - | Reduced (c,rest) -> (EConstr.of_constr (nf_beta sigma c), rest), nocase) + | Reduced (c,rest) -> (nf_beta sigma c, rest), nocase) | NotAnElimination when unfold_nonelim -> let c = reference_value env sigma ref u in - (EConstr.of_constr (whd_betaiotazeta sigma (applist (c, largs))), []), nocase + (whd_betaiotazeta sigma (applist (c, largs)), []), nocase | _ -> raise Redelimination with Redelimination when unfold_anyway -> let c = reference_value env sigma ref u in - (EConstr.of_constr (whd_betaiotazeta sigma (applist (c, largs))), []), nocase + (whd_betaiotazeta sigma (applist (c, largs)), []), nocase and reduce_params env sigma stack l = let len = List.length stack in @@ -721,7 +721,7 @@ and reduce_params env sigma stack l = let arg = List.nth stack i in let rarg = whd_construct_stack env sigma arg in match EConstr.kind sigma (fst rarg) with - | Construct _ -> List.assign stack i (EConstr.applist rarg) + | Construct _ -> List.assign stack i (applist rarg) | _ -> raise Redelimination) stack l @@ -817,9 +817,9 @@ and whd_construct_stack env sigma s = *) let try_red_product env sigma c = - let simpfun c = EConstr.of_constr (clos_norm_flags betaiotazeta env sigma c) in + let simpfun c = clos_norm_flags betaiotazeta env sigma c in let rec redrec env x = - let x = EConstr.of_constr (whd_betaiota sigma x) in + let x = whd_betaiota sigma x in match EConstr.kind sigma x with | App (f,l) -> (match EConstr.kind sigma f with @@ -856,7 +856,7 @@ let try_red_product env sigma c = | None -> raise Redelimination | Some c -> c) | _ -> raise Redelimination) - in EConstr.Unsafe.to_constr (redrec env c) + in redrec env c let red_product env sigma c = try try_red_product env sigma c @@ -953,7 +953,7 @@ let hnf_constr = whd_simpl_orelse_delta_but_fix (* The "simpl" reduction tactic *) let whd_simpl env sigma c = - EConstr.Unsafe.to_constr (EConstr.applist (whd_simpl_stack env sigma c)) + applist (whd_simpl_stack env sigma c) let simpl env sigma c = strong whd_simpl env sigma c @@ -1010,7 +1010,7 @@ let e_contextually byhead (occs,c) f = { e_redfun = begin fun env sigma t -> if locs != [] then ignore (traverse_below (Some (!pos-1)) envc t); let Sigma (t, evm, _) = (f subst).e_redfun env (Sigma.Unsafe.of_evar_map !evd) t in - (evd := Sigma.to_evar_map evm; EConstr.of_constr t) + (evd := Sigma.to_evar_map evm; t) end else traverse_below nested envc t @@ -1029,7 +1029,7 @@ let e_contextually byhead (occs,c) f = { e_redfun = begin fun env sigma t -> in let t' = traverse None (env,c) t in if List.exists (fun o -> o >= !pos) locs then error_invalid_occurrence locs; - Sigma.Unsafe.of_pair (EConstr.Unsafe.to_constr t', !evd) + Sigma.Unsafe.of_pair (t', !evd) end } let contextually byhead occs f env sigma t = @@ -1080,7 +1080,7 @@ let string_of_evaluable_ref env = function let unfold env sigma name c = if is_evaluable env name then - EConstr.of_constr (clos_norm_flags (unfold_red name) env sigma c) + clos_norm_flags (unfold_red name) env sigma c else error (string_of_evaluable_ref env name^" is opaque.") @@ -1098,7 +1098,7 @@ let unfoldoccs env sigma (occs,name) c = | [] -> () | _ -> error_invalid_occurrence rest in - EConstr.of_constr (nf_betaiotazeta sigma uc) + nf_betaiotazeta sigma uc in match occs with | NoOccurrences -> c @@ -1108,17 +1108,17 @@ let unfoldoccs env sigma (occs,name) c = (* Unfold reduction tactic: *) let unfoldn loccname env sigma c = - EConstr.Unsafe.to_constr (List.fold_left (fun c occname -> unfoldoccs env sigma occname c) c loccname) + List.fold_left (fun c occname -> unfoldoccs env sigma occname c) c loccname (* Re-folding constants tactics: refold com in term c *) let fold_one_com com env sigma c = let rcom = - try EConstr.of_constr (red_product env sigma com) + try red_product env sigma com with Redelimination -> error "Not reducible." in (* Reason first on the beta-iota-zeta normal form of the constant as unfold produces it, so that the "unfold f; fold f" configuration works to refold fix expressions *) - let a = subst_term sigma (EConstr.of_constr (clos_norm_flags unfold_side_red env sigma rcom)) c in + let a = subst_term sigma (clos_norm_flags unfold_side_red env sigma rcom) c in if not (EConstr.eq_constr sigma a c) then Vars.subst1 com a else @@ -1128,12 +1128,12 @@ let fold_one_com com env sigma c = Vars.subst1 com a let fold_commands cl env sigma c = - EConstr.Unsafe.to_constr (List.fold_right (fun com c -> fold_one_com com env sigma c) (List.rev cl) c) + List.fold_right (fun com c -> fold_one_com com env sigma c) (List.rev cl) c (* call by value reduction functions *) let cbv_norm_flags flags env sigma t = - cbv_norm (create_cbv_infos flags env sigma) t + EConstr.of_constr (cbv_norm (create_cbv_infos flags env sigma) t) let cbv_beta = cbv_norm_flags beta empty_env let cbv_betaiota = cbv_norm_flags betaiota empty_env @@ -1163,7 +1163,7 @@ let pattern_occs loccs_trm = { e_redfun = begin fun env sigma c -> let abstr_trm, sigma = List.fold_right (abstract_scheme env sigma) loccs_trm (c,sigma) in try let _ = Typing.unsafe_type_of env sigma abstr_trm in - Sigma.Unsafe.of_pair (EConstr.Unsafe.to_constr (applist(abstr_trm, List.map snd loccs_trm)), sigma) + Sigma.Unsafe.of_pair (applist(abstr_trm, List.map snd loccs_trm), sigma) with Type_errors.TypeError (env',t) -> raise (ReductionTacticError (InvalidAbstraction (env,sigma,abstr_trm,(env',t)))) end } @@ -1189,7 +1189,6 @@ let check_not_primitive_record env ind = let reduce_to_ind_gen allow_product env sigma t = let rec elimrec env t l = let t = hnf_constr env sigma t in - let t = EConstr.of_constr t in match EConstr.kind sigma (fst (decompose_app_vect sigma t)) with | Ind ind-> (check_privacy env ind, it_mkProd_or_LetIn t l) | Prod (n,ty,t') -> @@ -1202,7 +1201,6 @@ let reduce_to_ind_gen allow_product env sigma t = (* Last chance: we allow to bypass the Opaque flag (as it was partially the case between V5.10 and V8.1 *) let t' = whd_all env sigma t in - let t' = EConstr.of_constr t' in match EConstr.kind sigma (fst (decompose_app_vect sigma t')) with | Ind ind-> (check_privacy env ind, it_mkProd_or_LetIn t' l) | _ -> user_err (str"Not an inductive product.") @@ -1285,7 +1283,6 @@ let reduce_to_ref_gen allow_product env sigma ref t = with Not_found -> try let t' = nf_betaiota sigma (one_step_reduce env sigma t) in - let t' = EConstr.of_constr t' in elimrec env t' l with NotStepReducible -> error_cannot_recognize ref in diff --git a/pretyping/typeclasses.ml b/pretyping/typeclasses.ml index f59a6dcd9..9ee34341b 100644 --- a/pretyping/typeclasses.ml +++ b/pretyping/typeclasses.ml @@ -291,6 +291,7 @@ let build_subclasses ~check env sigma glob pri = let instapp = Reductionops.whd_beta sigma (EConstr.of_constr (appvectc c (Context.Rel.to_extended_vect 0 rels))) in + let instapp = EConstr.Unsafe.to_constr instapp in let projargs = Array.of_list (args @ [instapp]) in let projs = List.map_filter (fun (n, b, proj) -> diff --git a/pretyping/typing.ml b/pretyping/typing.ml index 40ef2450a..f67e0bddc 100644 --- a/pretyping/typing.ml +++ b/pretyping/typing.ml @@ -53,7 +53,7 @@ let inductive_type_knowing_parameters env sigma (ind,u) jl = EConstr.of_constr (Inductive.type_of_inductive_knowing_parameters env (mspec,u) paramstyp) let e_type_judgment env evdref j = - match EConstr.kind !evdref (EConstr.of_constr (whd_all env !evdref j.uj_type)) with + match EConstr.kind !evdref (whd_all env !evdref j.uj_type) with | Sort s -> {utj_val = j.uj_val; utj_type = s } | Evar ev -> let (evd,s) = Evardefine.define_evar_as_sort env !evdref ev in @@ -71,7 +71,7 @@ let e_judge_of_apply env evdref funj argjv = { uj_val = mkApp (j_val funj, Array.map j_val argjv); uj_type = typ } | hj::restjl -> - match EConstr.kind !evdref (EConstr.of_constr (whd_all env !evdref typ)) with + match EConstr.kind !evdref (whd_all env !evdref typ) with | Prod (_,c1,c2) -> if Evarconv.e_cumul env evdref hj.uj_type c1 then apply_rec (n+1) (subst1 hj.uj_val c2) restjl @@ -104,7 +104,7 @@ let e_is_correct_arity env evdref c pj ind specif params = let allowed_sorts = elim_sorts specif in let error () = Pretype_errors.error_elim_arity env !evdref ind allowed_sorts c pj None in let rec srec env pt ar = - let pt' = EConstr.of_constr (whd_all env !evdref pt) in + let pt' = whd_all env !evdref pt in match EConstr.kind !evdref pt', ar with | Prod (na1,a1,t), (LocalAssum (_,a1'))::ar' -> if not (Evarconv.e_cumul env evdref a1 (EConstr.of_constr a1')) then error (); @@ -144,7 +144,6 @@ let e_type_case_branches env evdref (ind,largs) pj c = let lc = Array.map EConstr.of_constr lc in let n = (snd specif).Declarations.mind_nrealdecls in let ty = whd_betaiota !evdref (lambda_applist_assum !evdref (n+1) p (realargs@[c])) in - let ty = EConstr.of_constr ty in (lc, ty) let e_judge_of_case env evdref ci pj cj lfj = diff --git a/pretyping/unification.ml b/pretyping/unification.ml index 169dd45bc..8a8649f11 100644 --- a/pretyping/unification.ml +++ b/pretyping/unification.ml @@ -75,7 +75,6 @@ let occur_meta_evd sigma mv c = let rec occrec c = (* Note: evars are not instantiated by terms with metas *) let c = whd_meta sigma c in - let c = EConstr.of_constr c in match EConstr.kind sigma c with | Meta mv' when Int.equal mv mv' -> raise Occur | _ -> EConstr.iter sigma occrec c @@ -1003,24 +1002,24 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst : subst0) conv_at_top e (match expand_key flags.modulo_delta curenv sigma cf1 with | Some c -> unirec_rec curenvnb pb opt substn - (EConstr.of_constr (whd_betaiotazeta sigma (mkApp(c,l1)))) cN + (whd_betaiotazeta sigma (mkApp(c,l1))) cN | None -> (match expand_key flags.modulo_delta curenv sigma cf2 with | Some c -> unirec_rec curenvnb pb opt substn cM - (EConstr.of_constr (whd_betaiotazeta sigma (mkApp(c,l2)))) + (whd_betaiotazeta sigma (mkApp(c,l2))) | None -> error_cannot_unify curenv sigma (cM,cN))) | Some false -> (match expand_key flags.modulo_delta curenv sigma cf2 with | Some c -> unirec_rec curenvnb pb opt substn cM - (EConstr.of_constr (whd_betaiotazeta sigma (mkApp(c,l2)))) + (whd_betaiotazeta sigma (mkApp(c,l2))) | None -> (match expand_key flags.modulo_delta curenv sigma cf1 with | Some c -> unirec_rec curenvnb pb opt substn - (EConstr.of_constr (whd_betaiotazeta sigma (mkApp(c,l1)))) cN + (whd_betaiotazeta sigma (mkApp(c,l1))) cN | None -> error_cannot_unify curenv sigma (cM,cN))) @@ -1233,7 +1232,7 @@ let applyHead env (type r) (evd : r Sigma.t) n c = Sigma (c, evd, p) else let sigma = Sigma.to_evar_map evd in - match EConstr.kind sigma (EConstr.of_constr (whd_all env sigma cty)) with + match EConstr.kind sigma (whd_all env sigma cty) with | Prod (_,c1,c2) -> let Sigma (evar, evd', q) = Evarutil.new_evar env evd ~src:(Loc.ghost,Evar_kinds.GoalEvar) c1 in apprec (n-1) (mkApp(c,[|evar|])) (subst1 evar c2) (p +> q) evd' @@ -1263,7 +1262,7 @@ let w_coerce_to_type env evd c cty mvty = but there are cases where it though it was not rigid (like in fst (nat,nat)) and stops while it could have seen that it is rigid *) let cty = Tacred.hnf_constr env evd cty in - try_to_coerce env evd c (EConstr.of_constr cty) tycon + try_to_coerce env evd c cty tycon let w_coerce env evd mv c = let cty = get_type_of env evd c in @@ -1276,7 +1275,6 @@ let unify_to_type env sigma flags c status u = let t = get_type_of env sigma (nf_meta sigma c) in let t = EConstr.of_constr t in let t = nf_betaiota sigma (nf_meta sigma t) in - let t = EConstr.of_constr t in unify_0 env sigma CUMUL flags t u let unify_type env sigma flags mv status c = @@ -1379,7 +1377,7 @@ let w_merge env with_types flags (evd,metas,evars : subst0) = else let evd' = if occur_meta_evd evd mv c then - if isMetaOf mv (whd_all env evd c) then evd + if isMetaOf evd mv (whd_all env evd c) then evd else error_cannot_unify env evd (mkMeta mv,c) else meta_assign mv (EConstr.Unsafe.to_constr c,(status,TypeProcessed)) evd in @@ -1618,7 +1616,7 @@ let make_pattern_test from_prefix_of_ind is_correct_type env sigma (pending,c) = (fun test -> match test.testing_state with | None -> None | Some (sigma,_,l) -> - let c = applist (EConstr.of_constr (nf_evar sigma (local_strong whd_meta sigma c)), l) in + let c = applist (EConstr.of_constr (nf_evar sigma (EConstr.Unsafe.to_constr (local_strong whd_meta sigma c))), l) in let univs, subst = nf_univ_variables sigma in Some (sigma,EConstr.of_constr (CVars.subst_univs_constr subst (EConstr.Unsafe.to_constr c)))) @@ -1877,7 +1875,6 @@ let w_unify_to_subterm_list env evd flags hdmeta oplist t = List.fold_right (fun op (evd,l) -> let op = whd_meta evd op in - let op = EConstr.of_constr op in if isMeta evd op then if flags.allow_K_in_toplevel_higher_order_unification then (evd,op::l) else error_abstraction_over_meta env evd hdmeta (destMeta evd op) @@ -1983,8 +1980,8 @@ let w_unify2 env evd flags dep cv_pb ty1 ty2 = convertible and first-order otherwise. But if failed if e.g. the type of Meta(1) had meta-variables in it. *) let w_unify env evd cv_pb ?(flags=default_unify_flags ()) ty1 ty2 = - let hd1,l1 = decompose_app_vect evd (EConstr.of_constr (whd_nored evd ty1)) in - let hd2,l2 = decompose_app_vect evd (EConstr.of_constr (whd_nored evd ty2)) in + let hd1,l1 = decompose_app_vect evd (whd_nored evd ty1) in + let hd2,l2 = decompose_app_vect evd (whd_nored evd ty2) in let is_empty1 = Array.is_empty l1 in let is_empty2 = Array.is_empty l2 in match EConstr.kind evd hd1, not is_empty1, EConstr.kind evd hd2, not is_empty2 with diff --git a/pretyping/vnorm.ml b/pretyping/vnorm.ml index 31693d82f..74998349b 100644 --- a/pretyping/vnorm.ml +++ b/pretyping/vnorm.ml @@ -360,7 +360,7 @@ let cbv_vm env sigma c t = let c = EConstr.to_constr sigma c in let t = EConstr.to_constr sigma t in let v = Vconv.val_of_constr env c in - nf_val env sigma v t + EConstr.of_constr (nf_val env sigma v t) let vm_infer_conv ?(pb=Reduction.CUMUL) env sigma t1 t2 = Reductionops.infer_conv_gen (fun pb ~l2r sigma ts -> Vconv.vm_conv_gen pb) diff --git a/pretyping/vnorm.mli b/pretyping/vnorm.mli index 650f3f291..8a4202c88 100644 --- a/pretyping/vnorm.mli +++ b/pretyping/vnorm.mli @@ -10,4 +10,4 @@ open EConstr open Environ (** {6 Reduction functions } *) -val cbv_vm : env -> Evd.evar_map -> constr -> types -> Constr.t +val cbv_vm : env -> Evd.evar_map -> constr -> types -> constr diff --git a/printing/prettyp.ml b/printing/prettyp.ml index ab0ce7e56..e66d3aafe 100644 --- a/printing/prettyp.ml +++ b/printing/prettyp.ml @@ -73,7 +73,8 @@ let print_ref reduce ref = let typ = Global.type_of_global_unsafe ref in let typ = if reduce then - let ctx,ccl = Reductionops.splay_prod_assum (Global.env()) Evd.empty (EConstr.of_constr typ) + let ctx,ccl = Reductionops.splay_prod_assum (Global.env()) Evd.empty (EConstr.of_constr typ) in + let ccl = EConstr.Unsafe.to_constr ccl in it_mkProd_or_LetIn ccl ctx else typ in let univs = Global.universes_of_global ref in @@ -595,6 +596,7 @@ let gallina_print_context with_values = let gallina_print_eval red_fun env sigma _ {uj_val=trm;uj_type=typ} = let ntrm = red_fun env sigma (EConstr.of_constr trm) in + let ntrm = EConstr.Unsafe.to_constr ntrm in (str " = " ++ gallina_print_typed_value_in_env env sigma (ntrm,typ)) (******************************************) diff --git a/proofs/clenv.ml b/proofs/clenv.ml index 572db1d40..a428ab0ed 100644 --- a/proofs/clenv.ml +++ b/proofs/clenv.ml @@ -60,7 +60,7 @@ let refresh_undefined_univs clenv = { clenv with evd = evd'; templval = map_freelisted clenv.templval; templtyp = map_freelisted clenv.templtyp }, subst -let clenv_hnf_constr ce t = EConstr.of_constr (hnf_constr (cl_env ce) (cl_sigma ce) t) +let clenv_hnf_constr ce t = hnf_constr (cl_env ce) (cl_sigma ce) t let clenv_get_type_of ce c = EConstr.of_constr (Retyping.get_type_of (cl_env ce) (cl_sigma ce) c) @@ -71,7 +71,6 @@ let mk_freelisted c = let clenv_push_prod cl = let typ = whd_all (cl_env cl) (cl_sigma cl) (clenv_type cl) in - let typ = EConstr.of_constr typ in let rec clrec typ = match EConstr.kind cl.evd typ with | Cast (t,_,_) -> clrec t | Prod (na,t,u) -> @@ -266,7 +265,7 @@ let clenv_unify_meta_types ?(flags=default_unify_flags ()) clenv = let clenv_unique_resolver ?(flags=default_unify_flags ()) clenv gl = let concl = Goal.V82.concl clenv.evd (sig_it gl) in - if isMeta clenv.evd (fst (decompose_app_vect clenv.evd (EConstr.of_constr (whd_nored clenv.evd clenv.templtyp.rebus)))) then + if isMeta clenv.evd (fst (decompose_app_vect clenv.evd (whd_nored clenv.evd clenv.templtyp.rebus))) then clenv_unify CUMUL ~flags (clenv_type clenv) concl (clenv_unify_meta_types ~flags clenv) else @@ -480,7 +479,7 @@ let error_already_defined b = (str "Position " ++ int n ++ str" already defined.") let clenv_unify_binding_type clenv c t u = - if isMeta clenv.evd (fst (decompose_app_vect clenv.evd (EConstr.of_constr (whd_nored clenv.evd u)))) then + if isMeta clenv.evd (fst (decompose_app_vect clenv.evd (whd_nored clenv.evd u))) then (* Not enough information to know if some subtyping is needed *) CoerceToType, clenv, c else @@ -498,7 +497,6 @@ let clenv_unify_binding_type clenv c t u = let clenv_assign_binding clenv k c = let k_typ = clenv_hnf_constr clenv (clenv_meta_type clenv k) in let c_typ = nf_betaiota clenv.evd (clenv_get_type_of clenv c) in - let c_typ = EConstr.of_constr c_typ in let status,clenv',c = clenv_unify_binding_type clenv c c_typ k_typ in let c = EConstr.Unsafe.to_constr c in { clenv' with evd = meta_assign k (c,(Conv,status)) clenv'.evd } diff --git a/proofs/logic.ml b/proofs/logic.ml index 829c96296..98ad9ebe3 100644 --- a/proofs/logic.ml +++ b/proofs/logic.ml @@ -34,7 +34,7 @@ type refiner_error = | CannotApply of constr * constr | NotWellTyped of constr | NonLinearProof of constr - | MetaInType of constr + | MetaInType of EConstr.constr (* Errors raised by the tactics *) | IntroNeedsProduct @@ -347,10 +347,11 @@ let rec mk_refgoals sigma goal goalacc conclty trm = match kind_of_term trm with | Meta _ -> let conclty = nf_betaiota sigma (EConstr.of_constr conclty) in - if !check && occur_meta sigma (EConstr.of_constr conclty) then + if !check && occur_meta sigma conclty then raise (RefinerError (MetaInType conclty)); - let (gl,ev,sigma) = mk_goal hyps (EConstr.of_constr conclty) in + let (gl,ev,sigma) = mk_goal hyps conclty in let ev = EConstr.Unsafe.to_constr ev in + let conclty = EConstr.Unsafe.to_constr conclty in gl::goalacc, conclty, sigma, ev | Cast (t,k, ty) -> @@ -425,7 +426,7 @@ and mk_hdgoals sigma goal goalacc trm = match kind_of_term trm with | Cast (c,_, ty) when isMeta c -> check_typability env sigma ty; - let (gl,ev,sigma) = mk_goal hyps (EConstr.of_constr (nf_betaiota sigma (EConstr.of_constr ty))) in + let (gl,ev,sigma) = mk_goal hyps (nf_betaiota sigma (EConstr.of_constr ty)) in let ev = EConstr.Unsafe.to_constr ev in gl::goalacc,ty,sigma,ev @@ -474,6 +475,7 @@ and mk_hdgoals sigma goal goalacc trm = and mk_arggoals sigma goal goalacc funty allargs = let foldmap (goalacc, funty, sigma) harg = let t = whd_all (Goal.V82.env sigma goal) sigma (EConstr.of_constr funty) in + let t = EConstr.Unsafe.to_constr t in let rec collapse t = match kind_of_term t with | LetIn (_, c1, _, b) -> collapse (subst1 c1 b) | _ -> t @@ -537,7 +539,7 @@ let prim_refiner r sigma goal = (* Logical rules *) | Cut (b,replace,id,t) -> (* if !check && not (Retyping.get_sort_of env sigma t) then*) - let (sg1,ev1,sigma) = mk_goal sign (EConstr.of_constr (nf_betaiota sigma (EConstr.of_constr t))) in + let (sg1,ev1,sigma) = mk_goal sign (nf_betaiota sigma (EConstr.of_constr t)) in let sign,t,cl,sigma = if replace then let nexthyp = get_hyp_after id (named_context_of_val sign) in diff --git a/proofs/logic.mli b/proofs/logic.mli index 8c8a01cad..5fe28ac76 100644 --- a/proofs/logic.mli +++ b/proofs/logic.mli @@ -43,7 +43,7 @@ type refiner_error = | CannotApply of constr * constr | NotWellTyped of constr | NonLinearProof of constr - | MetaInType of constr + | MetaInType of EConstr.constr (*i Errors raised by the tactics i*) | IntroNeedsProduct diff --git a/proofs/redexpr.ml b/proofs/redexpr.ml index d4a58da32..a830e25d9 100644 --- a/proofs/redexpr.ml +++ b/proofs/redexpr.ml @@ -44,7 +44,7 @@ let whd_cbn flags env sigma t = let (state,_) = (whd_state_gen true true flags env sigma (t,Reductionops.Stack.empty)) in - EConstr.Unsafe.to_constr (Reductionops.Stack.zip ~refold:true sigma state) + Reductionops.Stack.zip ~refold:true sigma state let strong_cbn flags = strong (whd_cbn flags) diff --git a/proofs/tacmach.ml b/proofs/tacmach.ml index f3f19e854..7ecf0a9e8 100644 --- a/proofs/tacmach.ml +++ b/proofs/tacmach.ml @@ -23,8 +23,6 @@ open Context.Named.Declaration module NamedDecl = Context.Named.Declaration -let compute env sigma c = EConstr.of_constr (compute env sigma c) - let re_sig it gc = { it = it; sigma = gc; } (**************************************************************) @@ -229,13 +227,13 @@ module New = struct let pf_reduce_to_quantified_ind gl t = pf_apply reduce_to_quantified_ind gl t - let pf_hnf_constr gl t = EConstr.of_constr (pf_apply hnf_constr gl t) + let pf_hnf_constr gl t = pf_apply hnf_constr gl t let pf_hnf_type_of gl t = - EConstr.of_constr (pf_whd_all gl (EConstr.of_constr (pf_get_type_of gl t))) + pf_whd_all gl (EConstr.of_constr (pf_get_type_of gl t)) let pf_matches gl pat t = pf_apply Constr_matching.matches_conv gl pat t - let pf_whd_all gl t = EConstr.of_constr (pf_apply whd_all gl t) + let pf_whd_all gl t = pf_apply whd_all gl t let pf_compute gl t = pf_apply compute gl t let pf_nf_evar gl t = nf_evar (project gl) t diff --git a/proofs/tacmach.mli b/proofs/tacmach.mli index 7b387ac9a..21511b6f9 100644 --- a/proofs/tacmach.mli +++ b/proofs/tacmach.mli @@ -42,7 +42,7 @@ val pf_ids_of_hyps : goal sigma -> Id.t list val pf_global : goal sigma -> Id.t -> constr val pf_unsafe_type_of : goal sigma -> EConstr.constr -> types val pf_type_of : goal sigma -> EConstr.constr -> evar_map * types -val pf_hnf_type_of : goal sigma -> EConstr.constr -> types +val pf_hnf_type_of : goal sigma -> EConstr.constr -> EConstr.types val pf_get_hyp : goal sigma -> Id.t -> Context.Named.Declaration.t val pf_get_hyp_typ : goal sigma -> Id.t -> types @@ -50,7 +50,7 @@ val pf_get_hyp_typ : goal sigma -> Id.t -> types val pf_get_new_id : Id.t -> goal sigma -> Id.t val pf_get_new_ids : Id.t list -> goal sigma -> Id.t list -val pf_reduction_of_red_expr : goal sigma -> red_expr -> EConstr.constr -> evar_map * constr +val pf_reduction_of_red_expr : goal sigma -> red_expr -> EConstr.constr -> evar_map * EConstr.constr val pf_apply : (env -> evar_map -> 'a) -> goal sigma -> 'a @@ -63,15 +63,15 @@ val pf_e_reduce : (env -> evar_map -> constr -> evar_map * constr) -> goal sigma -> constr -> evar_map * constr -val pf_whd_all : goal sigma -> EConstr.constr -> constr -val pf_hnf_constr : goal sigma -> EConstr.constr -> constr -val pf_nf : goal sigma -> EConstr.constr -> constr -val pf_nf_betaiota : goal sigma -> EConstr.constr -> constr +val pf_whd_all : goal sigma -> EConstr.constr -> EConstr.constr +val pf_hnf_constr : goal sigma -> EConstr.constr -> EConstr.constr +val pf_nf : goal sigma -> EConstr.constr -> EConstr.constr +val pf_nf_betaiota : goal sigma -> EConstr.constr -> EConstr.constr val pf_reduce_to_quantified_ind : goal sigma -> EConstr.types -> pinductive * EConstr.types val pf_reduce_to_atomic_ind : goal sigma -> EConstr.types -> pinductive * EConstr.types val pf_compute : goal sigma -> EConstr.constr -> EConstr.constr val pf_unfoldn : (occurrences * evaluable_global_reference) list - -> goal sigma -> EConstr.constr -> constr + -> goal sigma -> EConstr.constr -> EConstr.constr val pf_const_value : goal sigma -> pconstant -> constr val pf_conv_x : goal sigma -> constr -> constr -> bool diff --git a/tactics/autorewrite.ml b/tactics/autorewrite.ml index d656206d6..029384297 100644 --- a/tactics/autorewrite.ml +++ b/tactics/autorewrite.ml @@ -285,6 +285,7 @@ let decompose_applied_relation metas env sigma c ctype left2right = | Some c -> Some c | None -> let ctx,t' = Reductionops.splay_prod_assum env sigma (EConstr.of_constr ctype) in (* Search for underlying eq *) + let t' = EConstr.Unsafe.to_constr t' in match find_rel (it_mkProd_or_LetIn t' ctx) with | Some c -> Some c | None -> None diff --git a/tactics/class_tactics.ml b/tactics/class_tactics.ml index bf4e53b10..3a5347bbf 100644 --- a/tactics/class_tactics.ml +++ b/tactics/class_tactics.ml @@ -535,7 +535,6 @@ let make_resolve_hyp env sigma st flags only_classes pri decl = | _ -> let env' = Environ.push_rel_context ctx env in let ty' = Reductionops.whd_all env' sigma ar in - let ty' = EConstr.of_constr ty' in if not (EConstr.eq_constr sigma ty' ar) then iscl env' ty' else false in diff --git a/tactics/contradiction.ml b/tactics/contradiction.ml index 2d5c28eba..afc7e1547 100644 --- a/tactics/contradiction.ml +++ b/tactics/contradiction.ml @@ -69,7 +69,6 @@ let contradiction_context = let typ = nf_evar sigma (NamedDecl.get_type d) in let typ = EConstr.of_constr typ in let typ = whd_all env sigma typ in - let typ = EConstr.of_constr typ in if is_empty_type sigma typ then simplest_elim (mkVar id) else match EConstr.kind sigma typ with @@ -106,7 +105,7 @@ let contradiction_context = end } let is_negation_of env sigma typ t = - match EConstr.kind sigma (EConstr.of_constr (whd_all env sigma t)) with + match EConstr.kind sigma (whd_all env sigma t) with | Prod (na,t,u) -> is_empty_type sigma u && is_conv_leq env sigma typ t | _ -> false diff --git a/tactics/eauto.ml b/tactics/eauto.ml index 57400d334..92e59c5ce 100644 --- a/tactics/eauto.ml +++ b/tactics/eauto.ml @@ -481,7 +481,7 @@ let unfold_head env sigma (ids, csts) c = true, EConstr.of_constr (Environ.constant_value_in env c) | App (f, args) -> (match aux f with - | true, f' -> true, EConstr.of_constr (Reductionops.whd_betaiota sigma (mkApp (f', args))) + | true, f' -> true, Reductionops.whd_betaiota sigma (mkApp (f', args)) | false, _ -> let done_, args' = Array.fold_left_i (fun i (done_, acc) arg -> diff --git a/tactics/eqschemes.ml b/tactics/eqschemes.ml index 57bac25b9..a8ea7446f 100644 --- a/tactics/eqschemes.ml +++ b/tactics/eqschemes.ml @@ -605,9 +605,9 @@ let fix_r2l_forward_rew_scheme (c, ctx') = (mkLambda_or_LetIn (RelDecl.map_constr (liftn (-1) 1) p) (mkLambda_or_LetIn (RelDecl.map_constr (liftn (-1) 2) hp) (mkLambda_or_LetIn (RelDecl.map_constr (lift 2) ind) - (Reductionops.whd_beta Evd.empty + (EConstr.Unsafe.to_constr (Reductionops.whd_beta Evd.empty (EConstr.of_constr (applist (c, - Context.Rel.to_extended_list 3 indargs @ [mkRel 1;mkRel 3;mkRel 2]))))))) + Context.Rel.to_extended_list 3 indargs @ [mkRel 1;mkRel 3;mkRel 2])))))))) in c', ctx' | _ -> anomaly (Pp.str "Ill-formed non-dependent left-to-right rewriting scheme") diff --git a/tactics/equality.ml b/tactics/equality.ml index 209c9eb66..494f36d7d 100644 --- a/tactics/equality.ml +++ b/tactics/equality.ml @@ -413,7 +413,7 @@ let type_of_clause cls gl = match cls with let leibniz_rewrite_ebindings_clause cls lft2rgt tac c t l with_evars frzevars dep_proof_ok hdcncl = Proofview.Goal.nf_s_enter { s_enter = begin fun gl -> let evd = Sigma.to_evar_map (Proofview.Goal.sigma gl) in - let isatomic = isProd evd (EConstr.of_constr (whd_zeta evd hdcncl)) in + let isatomic = isProd evd (whd_zeta evd hdcncl) in let dep_fun = if isatomic then dependent else dependent_no_evar in let type_of_cls = type_of_clause cls gl in let dep = dep_proof_ok && dep_fun evd c type_of_cls in @@ -453,7 +453,7 @@ let general_rewrite_ebindings_clause cls lft2rgt occs frzevars dep_proof_ok ?tac let env = Proofview.Goal.env gl in let ctype = get_type_of env sigma c in let ctype = EConstr.of_constr ctype in - let rels, t = decompose_prod_assum sigma (EConstr.of_constr (whd_betaiotazeta sigma ctype)) in + let rels, t = decompose_prod_assum sigma (whd_betaiotazeta sigma ctype) in match match_with_equality_type sigma t with | Some (hdcncl,args) -> (* Fast path: direct leibniz-like rewrite *) let lft2rgt = adjust_rewriting_direction args lft2rgt in @@ -470,7 +470,6 @@ let general_rewrite_ebindings_clause cls lft2rgt occs frzevars dep_proof_ok ?tac Proofview.tclEVARMAP >>= fun sigma -> let env' = push_rel_context rels env in let rels',t' = splay_prod_assum env' sigma t in (* Search for underlying eq *) - let t' = EConstr.of_constr t' in match match_with_equality_type sigma t' with | Some (hdcncl,args) -> let lft2rgt = adjust_rewriting_direction args lft2rgt in @@ -1051,7 +1050,7 @@ let onNegatedEquality with_evars tac = let sigma = Tacmach.New.project gl in let ccl = Proofview.Goal.concl gl in let env = Proofview.Goal.env gl in - match EConstr.kind sigma (EConstr.of_constr (hnf_constr env sigma ccl)) with + match EConstr.kind sigma (hnf_constr env sigma ccl) with | Prod (_,t,u) when is_empty_type sigma u -> tclTHEN introf (onLastHypId (fun id -> @@ -1133,7 +1132,6 @@ let make_tuple env sigma (rterm,rty) lind = let minimal_free_rels env sigma (c,cty) = let cty_rels = free_rels sigma cty in let cty' = simpl env sigma cty in - let cty' = EConstr.of_constr cty' in let rels' = free_rels sigma cty' in if Int.Set.subset cty_rels rels' then (cty,cty_rels) @@ -1380,7 +1378,6 @@ let inject_if_homogenous_dependent_pair ty = let simplify_args env sigma t = (* Quick hack to reduce in arguments of eq only *) - let simpl env sigma c = EConstr.of_constr (simpl env sigma c) in match decompose_app sigma t with | eq, [t;c1;c2] -> applist (eq,[t;simpl env sigma c1;simpl env sigma c2]) | eq, [t1;c1;t2;c2] -> applist (eq,[t1;simpl env sigma c1;t2;simpl env sigma c2]) @@ -1591,7 +1588,6 @@ let subst_tuple_term env sigma dep_pair1 dep_pair2 b = let expected_goal = beta_applist sigma (abst_B,List.map fst e2_list) in (* Simulate now the normalisation treatment made by Logic.mk_refgoals *) let expected_goal = nf_betaiota sigma expected_goal in - let expected_goal = EConstr.of_constr expected_goal in (* Retype to get universes right *) let sigma, expected_goal_ty = Typing.type_of env sigma expected_goal in let sigma, _ = Typing.type_of env sigma body in diff --git a/tactics/hints.ml b/tactics/hints.ml index 2b310033c..231695c35 100644 --- a/tactics/hints.ml +++ b/tactics/hints.ml @@ -768,7 +768,7 @@ let make_exact_entry env sigma pri poly ?(name=PathAny) (c, cty, ctx) = code = with_uid (Give_exact (c, cty, ctx)); }) let make_apply_entry env sigma (eapply,hnf,verbose) pri poly ?(name=PathAny) (c, cty, ctx) = - let cty = if hnf then EConstr.of_constr (hnf_constr env sigma cty) else cty in + let cty = if hnf then hnf_constr env sigma cty else cty in match EConstr.kind sigma cty with | Prod _ -> let sigma' = Evd.merge_context_set univ_flexible sigma ctx in @@ -921,7 +921,6 @@ let make_trivial env sigma poly ?(name=PathAny) r = let c,ctx = fresh_global_or_constr env sigma poly r in let sigma = Evd.merge_context_set univ_flexible sigma ctx in let t = hnf_constr env sigma (EConstr.of_constr (unsafe_type_of env sigma c)) in - let t = EConstr.of_constr t in let hd = head_constr sigma t in let ce = mk_clenv_from_env env sigma None (c,t) in (Some hd, { pri=1; diff --git a/tactics/leminv.ml b/tactics/leminv.ml index 609fa1be8..ef3bfc9d0 100644 --- a/tactics/leminv.ml +++ b/tactics/leminv.ml @@ -125,7 +125,7 @@ let max_prefix_sign lid sign = | id::l -> snd (max_rec (id, sign_prefix id sign) l) *) let rec add_prods_sign env sigma t = - match EConstr.kind sigma (EConstr.of_constr (whd_all env sigma t)) with + match EConstr.kind sigma (whd_all env sigma t) with | Prod (na,c1,b) -> let id = id_of_name_using_hdchar env (EConstr.Unsafe.to_constr t) na in let b'= subst1 (mkVar id) b in @@ -180,7 +180,7 @@ let compute_first_inversion_scheme env sigma ind sort dep_option = (pty,goal) in let npty = nf_all env sigma pty in - let extenv = push_named (LocalAssum (p,npty)) env in + let extenv = push_named (nlocal_assum (p,npty)) env in extenv, goal (* [inversion_scheme sign I] diff --git a/tactics/tactics.ml b/tactics/tactics.ml index dabe78b34..5ee29c089 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -582,7 +582,6 @@ let fix ido n = match ido with let rec check_is_mutcoind env sigma cl = let b = whd_all env sigma cl in - let b = EConstr.of_constr b in match EConstr.kind sigma b with | Prod (na, c1, b) -> check_is_mutcoind (push_rel (local_assum (na,c1)) env) sigma b @@ -634,11 +633,11 @@ let cofix ido = match ido with (* Reduction and conversion tactics *) (**************************************************************) -type tactic_reduction = env -> evar_map -> constr -> Constr.constr +type tactic_reduction = env -> evar_map -> constr -> constr let pf_reduce_decl redfun where decl gl = let open Context.Named.Declaration in - let redfun' c = EConstr.of_constr (Tacmach.New.pf_apply redfun gl c) in + let redfun' c = Tacmach.New.pf_apply redfun gl c in match decl with | LocalAssum (id,ty) -> let ty = EConstr.of_constr ty in @@ -722,7 +721,7 @@ let bind_red_expr_occurrences occs nbcl redexp = let reduct_in_concl (redfun,sty) = Proofview.Goal.nf_enter { enter = begin fun gl -> - convert_concl_no_check (EConstr.of_constr (Tacmach.New.pf_apply redfun gl (Tacmach.New.pf_concl gl))) sty + convert_concl_no_check (Tacmach.New.pf_apply redfun gl (Tacmach.New.pf_concl gl)) sty end } let reduct_in_hyp ?(check=false) redfun (id,where) = @@ -742,23 +741,25 @@ let reduct_option ?(check=false) redfun = function let pf_e_reduce_decl redfun where decl gl = let open Context.Named.Declaration in let sigma = Proofview.Goal.sigma gl in - let redfun sigma c = redfun.e_redfun (Tacmach.New.pf_env gl) sigma (EConstr.of_constr c) in + let redfun sigma c = redfun.e_redfun (Tacmach.New.pf_env gl) sigma c in match decl with | LocalAssum (id,ty) -> + let ty = EConstr.of_constr ty in if where == InHypValueOnly then user_err (pr_id id ++ str " has no value."); let Sigma (ty', sigma, p) = redfun sigma ty in - Sigma (LocalAssum (id, ty'), sigma, p) + Sigma (nlocal_assum (id, ty'), sigma, p) | LocalDef (id,b,ty) -> + let b = EConstr.of_constr b in + let ty = EConstr.of_constr ty in let Sigma (b', sigma, p) = if where != InHypTypeOnly then redfun sigma b else Sigma.here b sigma in let Sigma (ty', sigma, q) = if where != InHypValueOnly then redfun sigma ty else Sigma.here ty sigma in - Sigma (LocalDef (id, b', ty'), sigma, p +> q) + Sigma (nlocal_def (id, b', ty'), sigma, p +> q) let e_reduct_in_concl ~check (redfun, sty) = Proofview.Goal.nf_s_enter { s_enter = begin fun gl -> let sigma = Proofview.Goal.sigma gl in let Sigma (c', sigma, p) = redfun.e_redfun (Tacmach.New.pf_env gl) sigma (Tacmach.New.pf_concl gl) in - let c' = EConstr.of_constr c' in Sigma (convert_concl ~check c' sty, sigma, p) end } @@ -779,7 +780,6 @@ let e_change_in_concl (redfun,sty) = Proofview.Goal.s_enter { s_enter = begin fun gl -> let sigma = Proofview.Goal.sigma gl in let Sigma (c, sigma, p) = redfun.e_redfun (Proofview.Goal.env gl) sigma (Proofview.Goal.raw_concl gl) in - let c = EConstr.of_constr c in Sigma (convert_concl_no_check c sty, sigma, p) end } @@ -787,18 +787,21 @@ let e_pf_change_decl (redfun : bool -> e_reduction_function) where decl env sigm let open Context.Named.Declaration in match decl with | LocalAssum (id,ty) -> + let ty = EConstr.of_constr ty in if where == InHypValueOnly then user_err (pr_id id ++ str " has no value."); - let Sigma (ty', sigma, p) = (redfun false).e_redfun env sigma (EConstr.of_constr ty) in - Sigma (LocalAssum (id, ty'), sigma, p) + let Sigma (ty', sigma, p) = (redfun false).e_redfun env sigma ty in + Sigma (nlocal_assum (id, ty'), sigma, p) | LocalDef (id,b,ty) -> + let b = EConstr.of_constr b in + let ty = EConstr.of_constr ty in let Sigma (b', sigma, p) = - if where != InHypTypeOnly then (redfun true).e_redfun env sigma (EConstr.of_constr b) else Sigma.here b sigma + if where != InHypTypeOnly then (redfun true).e_redfun env sigma b else Sigma.here b sigma in let Sigma (ty', sigma, q) = - if where != InHypValueOnly then (redfun false).e_redfun env sigma (EConstr.of_constr ty) else Sigma.here ty sigma + if where != InHypValueOnly then (redfun false).e_redfun env sigma ty else Sigma.here ty sigma in - Sigma (LocalDef (id,b',ty'), sigma, p +> q) + Sigma (nlocal_def (id,b',ty'), sigma, p +> q) let e_change_in_hyp redfun (id,where) = Proofview.Goal.s_enter { s_enter = begin fun gl -> @@ -818,21 +821,22 @@ let check_types env sigma mayneedglobalcheck deep newc origc = let t1 = EConstr.of_constr t1 in if deep then begin let t2 = Retyping.get_type_of env sigma origc in + let t2 = EConstr.of_constr t2 in let sigma, t2 = Evarsolve.refresh_universes - ~onlyalg:true (Some false) env sigma (EConstr.of_constr t2) in + ~onlyalg:true (Some false) env sigma t2 in let t2 = EConstr.of_constr t2 in let sigma, b = infer_conv ~pb:Reduction.CUMUL env sigma t1 t2 in if not b then if - isSort sigma (EConstr.of_constr (whd_all env sigma t1)) && - isSort sigma (EConstr.of_constr (whd_all env sigma t2)) + isSort sigma (whd_all env sigma t1) && + isSort sigma (whd_all env sigma t2) then (mayneedglobalcheck := true; sigma) else user_err ~hdr:"convert-check-hyp" (str "Types are incompatible.") else sigma end else - if not (isSort sigma (EConstr.of_constr (whd_all env sigma t1))) then + if not (isSort sigma (whd_all env sigma t1)) then user_err ~hdr:"convert-check-hyp" (str "Not a type.") else sigma @@ -843,7 +847,7 @@ let change_and_check cv_pb mayneedglobalcheck deep t = { e_redfun = begin fun en let sigma = check_types env sigma mayneedglobalcheck deep t' c in let sigma, b = infer_conv ~pb:cv_pb env sigma t' c in if not b then user_err ~hdr:"convert-check-hyp" (str "Not convertible."); - Sigma.Unsafe.of_pair (EConstr.Unsafe.to_constr t', sigma) + Sigma.Unsafe.of_pair (t', sigma) end } (* Use cumulativity only if changing the conclusion not a subterm *) @@ -858,7 +862,7 @@ let change_on_subterm cv_pb deep t where = { e_redfun = begin fun env sigma c -> env sigma c in if !mayneedglobalcheck then begin - try ignore (Typing.unsafe_type_of env (Sigma.to_evar_map sigma) (EConstr.of_constr c)) + try ignore (Typing.unsafe_type_of env (Sigma.to_evar_map sigma) c) with e when catchable_exception e -> error "Replacement would lead to an ill-typed term." end; @@ -1101,8 +1105,8 @@ let intros_replacing ids = (* User-level introduction tactics *) let lookup_hypothesis_as_renamed env ccl = function - | AnonHyp n -> Detyping.lookup_index_as_renamed env ccl n - | NamedHyp id -> Detyping.lookup_name_as_displayed env ccl id + | AnonHyp n -> Detyping.lookup_index_as_renamed env (EConstr.Unsafe.to_constr ccl) n + | NamedHyp id -> Detyping.lookup_name_as_displayed env (EConstr.Unsafe.to_constr ccl) id let lookup_hypothesis_as_renamed_gen red h gl = let env = Proofview.Goal.env gl in @@ -1110,11 +1114,11 @@ let lookup_hypothesis_as_renamed_gen red h gl = match lookup_hypothesis_as_renamed env ccl h with | None when red -> let (redfun, _) = Redexpr.reduction_of_red_expr env (Red true) in - let Sigma (c, _, _) = redfun.e_redfun env (Proofview.Goal.sigma gl) (EConstr.of_constr ccl) in + let Sigma (c, _, _) = redfun.e_redfun env (Proofview.Goal.sigma gl) ccl in aux c | x -> x in - try aux (EConstr.to_constr (Tacmach.New.project gl) (Proofview.Goal.concl gl)) + try aux (Proofview.Goal.concl gl) with Redelimination -> None let is_quantified_hypothesis id gl = @@ -1261,7 +1265,6 @@ let cut c = let typ = Typing.unsafe_type_of env sigma c in let typ = EConstr.of_constr typ in let typ = whd_all env sigma typ in - let typ = EConstr.of_constr typ in match EConstr.kind sigma typ with | Sort _ -> true | _ -> false @@ -1270,7 +1273,7 @@ let cut c = if is_sort then let id = next_name_away_with_default "H" Anonymous (Tacmach.New.pf_ids_of_hyps gl) in (** Backward compat: normalize [c]. *) - let c = if normalize_cut then EConstr.of_constr (local_strong whd_betaiota sigma c) else c in + let c = if normalize_cut then local_strong whd_betaiota sigma c else c in Refine.refine ~unsafe:true { run = begin fun h -> let Sigma (f, h, p) = Evarutil.new_evar ~principal:true env h (mkArrow c (Vars.lift 1 concl)) in let Sigma (x, h, q) = Evarutil.new_evar env h c in @@ -1755,7 +1758,6 @@ let general_apply with_delta with_destruct with_evars clear_flag (loc,(c,lbind : let thm_ty0 = nf_betaiota sigma (EConstr.of_constr (Retyping.get_type_of env sigma c)) in let try_apply thm_ty nprod = try - let thm_ty = EConstr.of_constr thm_ty in let n = nb_prod_modulo_zeta sigma thm_ty - nprod in if n<0 then error "Applied theorem has not enough premisses."; let clause = make_clenv_binding_apply env sigma (Some n) (c,thm_ty) lbind in @@ -1766,7 +1768,7 @@ let general_apply with_delta with_destruct with_evars clear_flag (loc,(c,lbind : let rec try_red_apply thm_ty (exn0, info) = try (* Try to head-reduce the conclusion of the theorem *) - let red_thm = try_red_product env sigma (EConstr.of_constr thm_ty) in + let red_thm = try_red_product env sigma thm_ty in tclORELSEOPT (try_apply red_thm concl_nprod) (function (e, info) -> match e with @@ -1880,7 +1882,6 @@ let progress_with_clause flags innerclause clause = let apply_in_once_main flags innerclause env sigma (d,lbind) = let thm = nf_betaiota sigma (EConstr.of_constr (Retyping.get_type_of env sigma d)) in - let thm = EConstr.of_constr thm in let rec aux clause = try progress_with_clause flags innerclause clause with e when CErrors.noncritical e -> @@ -2183,7 +2184,6 @@ let apply_type newcl args = let store = Proofview.Goal.extra gl in Refine.refine { run = begin fun sigma -> let newcl = nf_betaiota (Sigma.to_evar_map sigma) newcl (* As in former Logic.refine *) in - let newcl = EConstr.of_constr newcl in let Sigma (ev, sigma, p) = Evarutil.new_evar env sigma ~principal:true ~store newcl in Sigma (applist (ev, args), sigma, p) @@ -2377,7 +2377,6 @@ let rewrite_hyp_then assert_style with_evars thin l2r id tac = let type_of = Tacmach.New.pf_unsafe_type_of gl in let whd_all = Tacmach.New.pf_apply whd_all gl in let t = whd_all (EConstr.of_constr (type_of (mkVar id))) in - let t = EConstr.of_constr t in let eqtac, thin = match match_with_equality_type sigma t with | Some (hdcncl,[_;lhs;rhs]) -> if l2r && isVar sigma lhs && not (occur_var env sigma (destVar sigma lhs) rhs) then @@ -3039,7 +3038,7 @@ let unfold_body x = let xval = EConstr.of_constr xval in Tacticals.New.afterHyp x begin fun aft -> let hl = List.fold_right (fun decl cl -> (NamedDecl.get_id decl, InHyp) :: cl) aft [] in - let rfun _ _ c = EConstr.Unsafe.to_constr (replace_vars [x, xval] c) in + let rfun _ _ c = replace_vars [x, xval] c in let reducth h = reduct_in_hyp rfun h in let reductc = reduct_in_concl (rfun, DEFAULTcast) in Tacticals.New.tclTHENLIST [Tacticals.New.tclMAP reducth hl; reductc] @@ -3692,7 +3691,6 @@ let abstract_args gl generalize_vars dep id defined f args = let aux (prod, ctx, ctxenv, c, args, eqs, refls, nongenvars, vars, env) arg = let name, ty, arity = let rel, c = Reductionops.splay_prod_n env !sigma 1 prod in - let c = EConstr.of_constr c in let decl = List.hd rel in RelDecl.get_name decl, RelDecl.get_type decl, c in @@ -3855,9 +3853,7 @@ let specialize_eqs id gl = let acc' = it_mkLambda_or_LetIn acc ctx'' in let ty' = Tacred.whd_simpl env !evars ty' and acc' = Tacred.whd_simpl env !evars acc' in - let acc' = EConstr.of_constr acc' in - let ty' = Evarutil.nf_evar !evars ty' in - let ty' = EConstr.of_constr ty' in + let ty' = EConstr.of_constr (Evarutil.nf_evar !evars (EConstr.Unsafe.to_constr ty')) in if worked then tclTHENFIRST (Tacmach.internal_cut true id ty') (Proofview.V82.of_tactic (exact_no_check ((* refresh_universes_strict *) acc'))) gl @@ -4368,7 +4364,7 @@ let use_bindings env sigma elim must_be_closed (c,lbind) typ = let (sigma, c) = pose_all_metas_as_evars env indclause.evd (clenv_value indclause) in Sigma.Unsafe.of_pair (c, sigma) with e when catchable_exception e -> - try find_clause (EConstr.of_constr (try_red_product env sigma typ)) + try find_clause (try_red_product env sigma typ) with Redelimination -> raise e in find_clause typ @@ -4390,7 +4386,7 @@ let check_enough_applied env sigma elim = | None -> (* No eliminator given *) fun u -> - let t,_ = decompose_app sigma (EConstr.of_constr (whd_all env sigma u)) in isInd sigma t + let t,_ = decompose_app sigma (whd_all env sigma u) in isInd sigma t | Some elimc -> let elimt = Retyping.get_type_of env sigma (fst elimc) in let elimt = EConstr.of_constr elimt in @@ -4716,7 +4712,7 @@ let maybe_betadeltaiota_concl allowred gl = if not allowred then concl else let env = Proofview.Goal.env gl in - EConstr.of_constr (whd_all env sigma concl) + whd_all env sigma concl let reflexivity_red allowred = Proofview.Goal.enter { enter = begin fun gl -> diff --git a/tactics/tactics.mli b/tactics/tactics.mli index 630c660a1..b0d9dcb1c 100644 --- a/tactics/tactics.mli +++ b/tactics/tactics.mli @@ -129,7 +129,7 @@ val exact_proof : Constrexpr.constr_expr -> unit Proofview.tactic (** {6 Reduction tactics. } *) -type tactic_reduction = env -> evar_map -> constr -> Constr.constr +type tactic_reduction = env -> evar_map -> constr -> constr type change_arg = patvar_map -> constr Sigma.run diff --git a/toplevel/command.ml b/toplevel/command.ml index 1b5c4f660..1093b7fdc 100644 --- a/toplevel/command.ml +++ b/toplevel/command.ml @@ -79,7 +79,7 @@ let red_constant_entry n ce sigma = function let redfun env sigma c = let sigma = Sigma.Unsafe.of_evar_map sigma in let Sigma (c, _, _) = redfun.e_redfun env sigma c in - c + EConstr.Unsafe.to_constr c in { ce with const_entry_body = Future.chain ~greedy:true ~pure:true proof_out (fun ((body,ctx),eff) -> (under_binders env sigma redfun n body,ctx),eff) } @@ -962,7 +962,7 @@ let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation = in try let ctx, ar = Reductionops.splay_prod_n env !evdref 2 (EConstr.of_constr relty) in - match ctx, kind_of_term ar with + match ctx, EConstr.kind !evdref ar with | [LocalAssum (_,t); LocalAssum (_,u)], Sort (Prop Null) when Reductionops.is_conv env !evdref (EConstr.of_constr t) (EConstr.of_constr u) -> t | _, _ -> error () diff --git a/toplevel/himsg.ml b/toplevel/himsg.ml index 6e63a71fd..7eb189ef5 100644 --- a/toplevel/himsg.ml +++ b/toplevel/himsg.ml @@ -87,7 +87,7 @@ let to_unsafe_judgment j = on_judgment EConstr.Unsafe.to_constr j let j_nf_betaiotaevar sigma j = { uj_val = Evarutil.nf_evar sigma j.uj_val; - uj_type = Reductionops.nf_betaiota sigma (EConstr.of_constr j.uj_type) } + uj_type = EConstr.Unsafe.to_constr (Reductionops.nf_betaiota sigma (EConstr.of_constr j.uj_type)) } let jv_nf_betaiotaevar sigma jl = Array.map (j_nf_betaiotaevar sigma) jl @@ -362,6 +362,7 @@ let explain_actual_type env sigma j t reason = let env = make_all_name_different env in let j = j_nf_betaiotaevar sigma j in let t = Reductionops.nf_betaiota sigma (EConstr.of_constr t) in + let t = EConstr.Unsafe.to_constr t in (** Actually print *) let pe = pr_ne_context_of (str "In environment") env sigma in let pc = pr_lconstr_env env sigma (Environ.j_val j) in @@ -381,6 +382,7 @@ let explain_cant_apply_bad_type env sigma (n,exptyp,actualtyp) rator randl = let randl = jv_nf_betaiotaevar sigma randl in let exptyp = Evarutil.nf_evar sigma exptyp in let actualtyp = Reductionops.nf_betaiota sigma (EConstr.of_constr actualtyp) in + let actualtyp = EConstr.Unsafe.to_constr actualtyp in let rator = Evarutil.j_nf_evar sigma rator in let rator = to_unsafe_judgment rator in let env = make_all_name_different env in @@ -1115,6 +1117,7 @@ let explain_non_linear_proof c = spc () ++ str "because a metavariable has several occurrences." let explain_meta_in_type c = + let c = EConstr.Unsafe.to_constr c in str "In refiner, a meta appears in the type " ++ brk(1,1) ++ pr_lconstr c ++ str " of another meta" diff --git a/toplevel/obligations.ml b/toplevel/obligations.ml index aac1d1ed4..2a0ebf402 100644 --- a/toplevel/obligations.ml +++ b/toplevel/obligations.ml @@ -264,7 +264,7 @@ let pperror cmd = CErrors.user_err ~hdr:"Program" cmd let error s = pperror (str s) let reduce c = - Reductionops.clos_norm_flags CClosure.betaiota (Global.env ()) Evd.empty (EConstr.of_constr c) + EConstr.Unsafe.to_constr (Reductionops.clos_norm_flags CClosure.betaiota (Global.env ()) Evd.empty (EConstr.of_constr c)) exception NoObligations of Id.t option @@ -536,6 +536,8 @@ let declare_mutual_definition l = let subs, typ = (subst_body true x) in let term = snd (Reductionops.splay_lam_n (Global.env ()) Evd.empty len (EConstr.of_constr subs)) in let typ = snd (Reductionops.splay_prod_n (Global.env ()) Evd.empty len (EConstr.of_constr typ)) in + let term = EConstr.Unsafe.to_constr term in + let typ = EConstr.Unsafe.to_constr typ in x.prg_reduce term, x.prg_reduce typ, x.prg_implicits) l) in (* let fixdefs = List.map reduce_fix fixdefs in *) diff --git a/toplevel/record.ml b/toplevel/record.ml index 929505716..96221b742 100644 --- a/toplevel/record.ml +++ b/toplevel/record.ml @@ -121,6 +121,7 @@ let typecheck_params_and_fields def id pl t ps nots fs = | CSort (_, Misctypes.GType []) -> true | _ -> false in let s = interp_type_evars env evars ~impls:empty_internalization_env t in let sred = Reductionops.whd_all env !evars (EConstr.of_constr s) in + let sred = EConstr.Unsafe.to_constr sred in (match kind_of_term sred with | Sort s' -> (if poly then diff --git a/toplevel/vernacentries.ml b/toplevel/vernacentries.ml index 773766d7e..25b639fb0 100644 --- a/toplevel/vernacentries.ml +++ b/toplevel/vernacentries.ml @@ -1588,7 +1588,7 @@ let vernac_check_may_eval redexp glopt rc = match redexp with | None -> let l = Evar.Set.union (Evd.evars_of_term j.Environ.uj_val) (Evd.evars_of_term j.Environ.uj_type) in - let j = { j with Environ.uj_type = Reductionops.nf_betaiota sigma' (EConstr.of_constr j.Environ.uj_type) } in + let j = { j with Environ.uj_type = EConstr.Unsafe.to_constr (Reductionops.nf_betaiota sigma' (EConstr.of_constr j.Environ.uj_type)) } in Feedback.msg_notice (print_judgment env sigma' j ++ pr_ne_evar_set (fnl () ++ str "where" ++ fnl ()) (mt ()) sigma' l ++ Printer.pr_universe_ctx sigma uctx) -- cgit v1.2.3 From a327ae04966aa1c7786ae32157516e934068ea89 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Mon, 21 Nov 2016 21:33:14 +0100 Subject: Quote API using EConstr. --- plugins/quote/g_quote.ml4 | 7 ++- plugins/quote/quote.ml | 119 +++++++++++++++++++++----------------- plugins/setoid_ring/g_newring.ml4 | 2 +- plugins/setoid_ring/newring.ml | 3 +- plugins/setoid_ring/newring.mli | 2 +- proofs/tacmach.ml | 5 ++ proofs/tacmach.mli | 4 +- 7 files changed, 81 insertions(+), 61 deletions(-) diff --git a/plugins/quote/g_quote.ml4 b/plugins/quote/g_quote.ml4 index e7e6ecef9..79c429615 100644 --- a/plugins/quote/g_quote.ml4 +++ b/plugins/quote/g_quote.ml4 @@ -22,7 +22,8 @@ let loc = Loc.ghost let cont = Id.of_string "cont" let x = Id.of_string "x" -let make_cont (k : Val.t) (c : Constr.t) = +let make_cont (k : Val.t) (c : EConstr.t) = + let c = EConstr.Unsafe.to_constr c in let c = Tacinterp.Value.of_constr c in let tac = TacCall (loc, ArgVar (loc, cont), [Reference (ArgVar (loc, x))]) in let ist = { lfun = Id.Map.add cont k (Id.Map.singleton x c); extra = TacStore.empty; } in @@ -32,8 +33,8 @@ TACTIC EXTEND quote [ "quote" ident(f) ] -> [ quote f [] ] | [ "quote" ident(f) "[" ne_ident_list(lc) "]"] -> [ quote f lc ] | [ "quote" ident(f) "in" constr(c) "using" tactic(k) ] -> - [ gen_quote (make_cont k) c f [] ] + [ gen_quote (make_cont k) (EConstr.of_constr c) f [] ] | [ "quote" ident(f) "[" ne_ident_list(lc) "]" "in" constr(c) "using" tactic(k) ] -> - [ gen_quote (make_cont k) c f lc ] + [ gen_quote (make_cont k) (EConstr.of_constr c) f lc ] END diff --git a/plugins/quote/quote.ml b/plugins/quote/quote.ml index 5f8a0b2d5..2ad97c75b 100644 --- a/plugins/quote/quote.ml +++ b/plugins/quote/quote.ml @@ -105,6 +105,7 @@ open CErrors open Util open Names open Term +open EConstr open Pattern open Patternops open Constr_matching @@ -116,7 +117,8 @@ open Proofview.Notations We do that lazily, because this code can be linked before the constants are loaded in the environment *) -let constant dir s = Coqlib.gen_constant "Quote" ("quote"::dir) s +let constant dir s = + EConstr.of_constr (Coqlib.gen_constant "Quote" ("quote"::dir) s) let coq_Empty_vm = lazy (constant ["Quote"] "Empty_vm") let coq_Node_vm = lazy (constant ["Quote"] "Node_vm") @@ -165,7 +167,7 @@ exchange ?1 and ?2 in the example above) module ConstrSet = Set.Make( struct - type t = constr + type t = Constr.constr let compare = constr_ord end) @@ -183,7 +185,7 @@ type inversion_scheme = { let i_can't_do_that () = error "Quote: not a simple fixpoint" -let decomp_term gl c = kind_of_term (EConstr.Unsafe.to_constr (Termops.strip_outer_cast (Tacmach.New.project gl) (EConstr.of_constr c))) +let decomp_term sigma c = EConstr.kind sigma (Termops.strip_outer_cast sigma c) (*s [compute_lhs typ i nargsi] builds the term \texttt{(C ?nargsi ... ?2 ?1)}, where \texttt{C} is the [i]-th constructor of inductive @@ -195,8 +197,8 @@ let coerce_meta_out id = let coerce_meta_in n = Id.of_string ("M" ^ string_of_int n) -let compute_lhs typ i nargsi = - match kind_of_term typ with +let compute_lhs sigma typ i nargsi = + match EConstr.kind sigma typ with | Ind((sp,0),u) -> let argsi = Array.init nargsi (fun j -> mkMeta (nargsi - j)) in mkApp (mkConstructU (((sp,0),i+1),u), argsi) @@ -205,60 +207,61 @@ let compute_lhs typ i nargsi = (*s This function builds the pattern from the RHS. Recursive calls are replaced by meta-variables ?i corresponding to those in the LHS *) -let compute_rhs bodyi index_of_f = +let compute_rhs env sigma bodyi index_of_f = let rec aux c = - match kind_of_term c with - | App (j, args) when isRel j && Int.equal (destRel j) index_of_f (* recursive call *) -> - let i = destRel (Array.last args) in + match EConstr.kind sigma c with + | App (j, args) when isRel sigma j && Int.equal (destRel sigma j) index_of_f (* recursive call *) -> + let i = destRel sigma (Array.last args) in PMeta (Some (coerce_meta_in i)) | App (f,args) -> - PApp (pattern_of_constr (Global.env()) Evd.empty (EConstr.of_constr f), Array.map aux args) + PApp (pattern_of_constr env sigma f, Array.map aux args) | Cast (c,_,_) -> aux c - | _ -> pattern_of_constr (Global.env())(*FIXME*) Evd.empty (EConstr.of_constr c) + | _ -> pattern_of_constr env sigma c in aux bodyi (*s Now the function [compute_ivs] itself *) let compute_ivs f cs gl = - let cst = try destConst f with DestKO -> i_can't_do_that () in + let env = Proofview.Goal.env gl in + let sigma = Tacmach.New.project gl in + let cst = try destConst sigma f with DestKO -> i_can't_do_that () in let body = Environ.constant_value_in (Global.env()) cst in - match decomp_term gl body with + let body = EConstr.of_constr body in + match decomp_term sigma body with | Fix(([| len |], 0), ([| name |], [| typ |], [| body2 |])) -> - let (args3, body3) = decompose_lam body2 in + let (args3, body3) = decompose_lam sigma body2 in let nargs3 = List.length args3 in - let env = Proofview.Goal.env gl in - let sigma = Tacmach.New.project gl in let is_conv = Reductionops.is_conv env sigma in - begin match decomp_term gl body3 with + begin match decomp_term sigma body3 with | Case(_,p,c,lci) -> (*

Case c of c1 ... cn end *) let n_lhs_rhs = ref [] and v_lhs = ref (None : constr option) and c_lhs = ref (None : constr option) in Array.iteri (fun i ci -> - let argsi, bodyi = decompose_lam ci in + let argsi, bodyi = decompose_lam sigma ci in let nargsi = List.length argsi in (* REL (narg3 + nargsi + 1) is f *) (* REL nargsi+1 to REL nargsi + nargs3 are arguments of f *) (* REL 1 to REL nargsi are argsi (reverse order) *) (* First we test if the RHS is the RHS for constants *) - if isRel bodyi && Int.equal (destRel bodyi) 1 then - c_lhs := Some (compute_lhs (snd (List.hd args3)) + if isRel sigma bodyi && Int.equal (destRel sigma bodyi) 1 then + c_lhs := Some (compute_lhs sigma (snd (List.hd args3)) i nargsi) (* Then we test if the RHS is the RHS for variables *) - else begin match decompose_app bodyi with + else begin match decompose_app sigma bodyi with | vmf, [_; _; a3; a4 ] - when isRel a3 && isRel a4 && is_conv (EConstr.of_constr vmf) - (EConstr.of_constr (Lazy.force coq_varmap_find)) -> - v_lhs := Some (compute_lhs + when isRel sigma a3 && isRel sigma a4 && is_conv vmf + (Lazy.force coq_varmap_find) -> + v_lhs := Some (compute_lhs sigma (snd (List.hd args3)) i nargsi) (* Third case: this is a normal LHS-RHS *) | _ -> n_lhs_rhs := - (compute_lhs (snd (List.hd args3)) i nargsi, - compute_rhs bodyi (nargs3 + nargsi + 1)) + (compute_lhs sigma (snd (List.hd args3)) i nargsi, + compute_rhs env sigma bodyi (nargs3 + nargsi + 1)) :: !n_lhs_rhs end) lci; @@ -266,8 +269,8 @@ let compute_ivs f cs gl = if Option.is_empty !c_lhs && Option.is_empty !v_lhs then i_can't_do_that (); (* The Cases predicate is a lambda; we assume no dependency *) - let p = match kind_of_term p with - | Lambda (_,_,p) -> Termops.pop (EConstr.of_constr p) + let p = match EConstr.kind sigma p with + | Lambda (_,_,p) -> EConstr.of_constr (Termops.pop p) | _ -> p in @@ -297,11 +300,11 @@ binary search trees (see file \texttt{Quote.v}) *) (* First the function to distinghish between constants (closed terms) and variables (open terms) *) -let rec closed_under cset t = - (ConstrSet.mem t cset) || - (match (kind_of_term t) with - | Cast(c,_,_) -> closed_under cset c - | App(f,l) -> closed_under cset f && Array.for_all (closed_under cset) l +let rec closed_under sigma cset t = + (ConstrSet.mem (EConstr.Unsafe.to_constr t) cset) || + (match EConstr.kind sigma t with + | Cast(c,_,_) -> closed_under sigma cset c + | App(f,l) -> closed_under sigma cset f && Array.for_all (closed_under sigma cset) l | _ -> false) (*s [btree_of_array [| c1; c2; c3; c4; c5 |]] builds the complete @@ -361,7 +364,7 @@ let path_of_int n = let rec subterm gl (t : constr) (t' : constr) = (pf_conv_x gl t t') || - (match (kind_of_term t) with + (match EConstr.kind (project gl) t with | App (f,args) -> Array.exists (fun t -> subterm gl t t') args | Cast(t,_,_) -> (subterm gl t t') | _ -> false) @@ -370,9 +373,10 @@ let rec subterm gl (t : constr) (t' : constr) = (* Since it's a partial order the algoritm of Sort.list won't work !! *) let rec sort_subterm gl l = + let sigma = project gl in let rec insert c = function | [] -> [c] - | (h::t as l) when eq_constr c h -> l (* Avoid doing the same work twice *) + | (h::t as l) when EConstr.eq_constr sigma c h -> l (* Avoid doing the same work twice *) | h::t -> if subterm gl c h then c::h::t else h::(insert c t) in match l with @@ -380,11 +384,15 @@ let rec sort_subterm gl l = | h::t -> insert h (sort_subterm gl t) module Constrhash = Hashtbl.Make - (struct type t = constr - let equal = eq_constr - let hash = hash_constr + (struct type t = Constr.constr + let equal = Term.eq_constr + let hash = Term.hash_constr end) +let subst_meta subst c = + let subst = List.map (fun (i, c) -> i, EConstr.Unsafe.to_constr c) subst in + EConstr.of_constr (Termops.subst_meta subst (EConstr.Unsafe.to_constr c)) + (*s Now we are able to do the inversion itself. We destructurate the term and use an imperative hashtable to store leafs that are already encountered. @@ -392,7 +400,7 @@ module Constrhash = Hashtbl.Make [ivs : inversion_scheme]\\ [lc: constr list]\\ [gl: goal sigma]\\ *) -let quote_terms ivs lc = +let quote_terms env sigma ivs lc = Coqlib.check_required_library ["Coq";"quote";"Quote"]; let varhash = (Constrhash.create 17 : constr Constrhash.t) in let varlist = ref ([] : constr list) in (* list of variables *) @@ -402,34 +410,34 @@ let quote_terms ivs lc = match l with | (lhs, rhs)::tail -> begin try - let s1 = Id.Map.bindings (matches (Global.env ()) Evd.empty rhs (EConstr.of_constr c)) in - let s2 = List.map (fun (i,c_i) -> (coerce_meta_out i,aux (EConstr.Unsafe.to_constr c_i))) s1 + let s1 = Id.Map.bindings (matches env sigma rhs c) in + let s2 = List.map (fun (i,c_i) -> (coerce_meta_out i,aux c_i)) s1 in - Termops.subst_meta s2 lhs + subst_meta s2 lhs with PatternMatchingFailure -> auxl tail end | [] -> begin match ivs.variable_lhs with | None -> begin match ivs.constant_lhs with - | Some c_lhs -> Termops.subst_meta [1, c] c_lhs + | Some c_lhs -> subst_meta [1, c] c_lhs | None -> anomaly (Pp.str "invalid inversion scheme for quote") end | Some var_lhs -> begin match ivs.constant_lhs with - | Some c_lhs when closed_under ivs.constants c -> - Termops.subst_meta [1, c] c_lhs + | Some c_lhs when closed_under sigma ivs.constants c -> + subst_meta [1, c] c_lhs | _ -> begin - try Constrhash.find varhash c + try Constrhash.find varhash (EConstr.Unsafe.to_constr c) with Not_found -> let newvar = - Termops.subst_meta [1, (path_of_int !counter)] + subst_meta [1, (path_of_int !counter)] var_lhs in begin incr counter; varlist := c :: !varlist; - Constrhash.add varhash c newvar; + Constrhash.add varhash (EConstr.Unsafe.to_constr c) newvar; newvar end end @@ -448,27 +456,32 @@ let quote_terms ivs lc = let quote f lid = Proofview.Goal.nf_enter { enter = begin fun gl -> + let env = Proofview.Goal.env gl in + let sigma = Tacmach.New.project gl in let f = Tacmach.New.pf_global f gl in + let f = EConstr.of_constr f in let cl = List.map (fun id -> Tacmach.New.pf_global id gl) lid in let ivs = compute_ivs f cl gl in let concl = Proofview.Goal.concl gl in - let concl = EConstr.Unsafe.to_constr concl in - let quoted_terms = quote_terms ivs [concl] in + let quoted_terms = quote_terms env sigma ivs [concl] in let (p, vm) = match quoted_terms with | [p], vm -> (p,vm) | _ -> assert false in match ivs.variable_lhs with - | None -> Tactics.convert_concl (EConstr.of_constr (mkApp (f, [| p |]))) DEFAULTcast - | Some _ -> Tactics.convert_concl (EConstr.of_constr (mkApp (f, [| vm; p |]))) DEFAULTcast + | None -> Tactics.convert_concl (mkApp (f, [| p |])) DEFAULTcast + | Some _ -> Tactics.convert_concl (mkApp (f, [| vm; p |])) DEFAULTcast end } let gen_quote cont c f lid = Proofview.Goal.nf_enter { enter = begin fun gl -> + let env = Proofview.Goal.env gl in + let sigma = Tacmach.New.project gl in let f = Tacmach.New.pf_global f gl in + let f = EConstr.of_constr f in let cl = List.map (fun id -> Tacmach.New.pf_global id gl) lid in let ivs = compute_ivs f cl gl in - let quoted_terms = quote_terms ivs [c] in + let quoted_terms = quote_terms env sigma ivs [c] in let (p, vm) = match quoted_terms with | [p], vm -> (p,vm) | _ -> assert false diff --git a/plugins/setoid_ring/g_newring.ml4 b/plugins/setoid_ring/g_newring.ml4 index 0987c44ae..89ec4c1c5 100644 --- a/plugins/setoid_ring/g_newring.ml4 +++ b/plugins/setoid_ring/g_newring.ml4 @@ -30,7 +30,7 @@ END TACTIC EXTEND closed_term [ "closed_term" constr(t) "[" ne_reference_list(l) "]" ] -> - [ closed_term t l ] + [ closed_term (EConstr.of_constr t) l ] END open Pptactic diff --git a/plugins/setoid_ring/newring.ml b/plugins/setoid_ring/newring.ml index b720b2e0a..2f6c00c9d 100644 --- a/plugins/setoid_ring/newring.ml +++ b/plugins/setoid_ring/newring.ml @@ -97,9 +97,10 @@ let protect_tac_in map id = let closed_term t l = let open Quote_plugin in + Proofview.tclEVARMAP >>= fun sigma -> let l = List.map Universes.constr_of_global l in let cs = List.fold_right Quote.ConstrSet.add l Quote.ConstrSet.empty in - if Quote.closed_under cs t then Proofview.tclUNIT () else Tacticals.New.tclFAIL 0 (mt()) + if Quote.closed_under sigma cs t then Proofview.tclUNIT () else Tacticals.New.tclFAIL 0 (mt()) (* TACTIC EXTEND echo | [ "echo" constr(t) ] -> diff --git a/plugins/setoid_ring/newring.mli b/plugins/setoid_ring/newring.mli index f417c87cd..89538eb24 100644 --- a/plugins/setoid_ring/newring.mli +++ b/plugins/setoid_ring/newring.mli @@ -19,7 +19,7 @@ val protect_tac_in : string -> Id.t -> unit Proofview.tactic val protect_tac : string -> unit Proofview.tactic -val closed_term : constr -> global_reference list -> unit Proofview.tactic +val closed_term : EConstr.constr -> global_reference list -> unit Proofview.tactic val process_ring_mods : constr_expr ring_mod list -> diff --git a/proofs/tacmach.ml b/proofs/tacmach.ml index 7ecf0a9e8..d9caadd54 100644 --- a/proofs/tacmach.ml +++ b/proofs/tacmach.ml @@ -41,6 +41,11 @@ let project = Refiner.project let pf_env = Refiner.pf_env let pf_hyps = Refiner.pf_hyps +let test_conversion cb env sigma c1 c2 = + let c1 = EConstr.Unsafe.to_constr c1 in + let c2 = EConstr.Unsafe.to_constr c2 in + test_conversion cb env sigma c1 c2 + let pf_concl gls = EConstr.Unsafe.to_constr (Goal.V82.concl (project gls) (sig_it gls)) let pf_hyps_types gls = let sign = Environ.named_context (pf_env gls) in diff --git a/proofs/tacmach.mli b/proofs/tacmach.mli index 21511b6f9..06b3e3981 100644 --- a/proofs/tacmach.mli +++ b/proofs/tacmach.mli @@ -74,8 +74,8 @@ val pf_unfoldn : (occurrences * evaluable_global_reference) list -> goal sigma -> EConstr.constr -> EConstr.constr val pf_const_value : goal sigma -> pconstant -> constr -val pf_conv_x : goal sigma -> constr -> constr -> bool -val pf_conv_x_leq : goal sigma -> constr -> constr -> bool +val pf_conv_x : goal sigma -> EConstr.constr -> EConstr.constr -> bool +val pf_conv_x_leq : goal sigma -> EConstr.constr -> EConstr.constr -> bool val pf_matches : goal sigma -> constr_pattern -> EConstr.constr -> patvar_map val pf_is_matching : goal sigma -> constr_pattern -> EConstr.constr -> bool -- cgit v1.2.3 From fa638c3e71752b6a59261776b36f1bed7d9c26d2 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Thu, 24 Nov 2016 12:07:55 +0100 Subject: Cc API using EConstr. --- plugins/cc/cctac.ml | 138 ++++++++++++++++++++++++-------------------- plugins/cc/cctac.mli | 2 +- plugins/cc/g_congruence.ml4 | 4 +- pretyping/reductionops.ml | 8 +++ pretyping/reductionops.mli | 1 + 5 files changed, 86 insertions(+), 67 deletions(-) diff --git a/plugins/cc/cctac.ml b/plugins/cc/cctac.ml index 130f01e97..0d48b65d0 100644 --- a/plugins/cc/cctac.ml +++ b/plugins/cc/cctac.ml @@ -13,6 +13,7 @@ open Names open Inductiveops open Declarations open Term +open EConstr open Vars open Tacmach open Tactics @@ -27,6 +28,10 @@ open Proofview.Notations module RelDecl = Context.Rel.Declaration module NamedDecl = Context.Named.Declaration +let local_assum (na, t) = + let inj = EConstr.Unsafe.to_constr in + RelDecl.LocalAssum (na, inj t) + let reference dir s = lazy (Coqlib.gen_reference "CC" dir s) let _f_equal = reference ["Init";"Logic"] "f_equal" @@ -39,27 +44,26 @@ let _False = reference ["Init";"Logic"] "False" let _True = reference ["Init";"Logic"] "True" let _I = reference ["Init";"Logic"] "I" -let whd env= - let infos=CClosure.create_clos_infos CClosure.betaiotazeta env in - (fun t -> CClosure.whd_val infos (CClosure.inject t)) +let whd env sigma t = + Reductionops.clos_whd_flags CClosure.betaiotazeta env sigma t -let whd_delta env= - let infos=CClosure.create_clos_infos CClosure.all env in - (fun t -> CClosure.whd_val infos (CClosure.inject t)) +let whd_delta env sigma t = + Reductionops.clos_whd_flags CClosure.all env sigma t (* decompose member of equality in an applicative format *) (** FIXME: evar leak *) -let sf_of env sigma c = e_sort_of env (ref sigma) (EConstr.of_constr c) +let sf_of env sigma c = e_sort_of env (ref sigma) c let rec decompose_term env sigma t= - match kind_of_term (whd env t) with + match EConstr.kind sigma (whd env sigma t) with App (f,args)-> let tf=decompose_term env sigma f in let targs=Array.map (decompose_term env sigma) args in Array.fold_left (fun s t->Appli (s,t)) tf targs - | Prod (_,a,_b) when EConstr.Vars.noccurn sigma 1 (EConstr.of_constr _b) -> - let b = Termops.pop (EConstr.of_constr _b) in + | Prod (_,a,_b) when noccurn sigma 1 _b -> + let b = Termops.pop _b in + let b = EConstr.of_constr b in let sort_b = sf_of env sigma b in let sort_a = sf_of env sigma a in Appli(Appli(Product (sort_a,sort_b) , @@ -77,28 +81,27 @@ let rec decompose_term env sigma t= | Ind c -> let (mind,i_ind),u = c in let canon_mind = mind_of_kn (canonical_mind mind) in - let canon_ind = canon_mind,i_ind in (Symb (mkIndU (canon_ind,u))) + let canon_ind = canon_mind,i_ind in (Symb (Constr.mkIndU (canon_ind,u))) | Const (c,u) -> let canon_const = constant_of_kn (canonical_con c) in - (Symb (mkConstU (canon_const,u))) + (Symb (Constr.mkConstU (canon_const,u))) | Proj (p, c) -> let canon_const kn = constant_of_kn (canonical_con kn) in let p' = Projection.map canon_const p in - (Appli (Symb (mkConst (Projection.constant p')), decompose_term env sigma c)) + (Appli (Symb (Constr.mkConst (Projection.constant p')), decompose_term env sigma c)) | _ -> - let t = Termops.strip_outer_cast sigma (EConstr.of_constr t) in - let t = EConstr.Unsafe.to_constr t in - if closed0 t then Symb t else raise Not_found + let t = Termops.strip_outer_cast sigma t in + if closed0 sigma t then Symb (EConstr.to_constr sigma t) else raise Not_found (* decompose equality in members and type *) -open Globnames +open Termops let atom_of_constr env sigma term = - let wh = (whd_delta env term) in - let kot = kind_of_term wh in + let wh = whd_delta env sigma term in + let kot = EConstr.kind sigma wh in match kot with App (f,args)-> - if is_global (Lazy.force _eq) f && Int.equal (Array.length args) 3 + if is_global sigma (Lazy.force _eq) f && Int.equal (Array.length args) 3 then `Eq (args.(0), decompose_term env sigma args.(1), decompose_term env sigma args.(2)) @@ -106,15 +109,16 @@ let atom_of_constr env sigma term = | _ -> `Other (decompose_term env sigma term) let rec pattern_of_constr env sigma c = - match kind_of_term (whd env c) with + match EConstr.kind sigma (whd env sigma c) with App (f,args)-> let pf = decompose_term env sigma f in let pargs,lrels = List.split (Array.map_to_list (pattern_of_constr env sigma) args) in PApp (pf,List.rev pargs), List.fold_left Int.Set.union Int.Set.empty lrels - | Prod (_,a,_b) when EConstr.Vars.noccurn sigma 1 (EConstr.of_constr _b) -> - let b = Termops.pop (EConstr.of_constr _b) in + | Prod (_,a,_b) when noccurn sigma 1 _b -> + let b = Termops.pop _b in + let b = EConstr.of_constr b in let pa,sa = pattern_of_constr env sigma a in let pb,sb = pattern_of_constr env sigma b in let sort_b = sf_of env sigma b in @@ -132,19 +136,19 @@ let non_trivial = function let patterns_of_constr env sigma nrels term= let f,args= - try destApp (whd_delta env term) with DestKO -> raise Not_found in - if is_global (Lazy.force _eq) f && Int.equal (Array.length args) 3 + try destApp sigma (whd_delta env sigma term) with DestKO -> raise Not_found in + if is_global sigma (Lazy.force _eq) f && Int.equal (Array.length args) 3 then let patt1,rels1 = pattern_of_constr env sigma args.(1) and patt2,rels2 = pattern_of_constr env sigma args.(2) in let valid1 = if not (Int.equal (Int.Set.cardinal rels1) nrels) then Creates_variables else if non_trivial patt1 then Normal - else Trivial args.(0) + else Trivial (EConstr.to_constr sigma args.(0)) and valid2 = if not (Int.equal (Int.Set.cardinal rels2) nrels) then Creates_variables else if non_trivial patt2 then Normal - else Trivial args.(0) in + else Trivial (EConstr.to_constr sigma args.(0)) in if valid1 != Creates_variables || valid2 != Creates_variables then nrels,valid1,patt1,valid2,patt2 @@ -152,28 +156,28 @@ let patterns_of_constr env sigma nrels term= else raise Not_found let rec quantified_atom_of_constr env sigma nrels term = - match kind_of_term (whd_delta env term) with + match EConstr.kind sigma (whd_delta env sigma term) with Prod (id,atom,ff) -> - if is_global (Lazy.force _False) ff then + if is_global sigma (Lazy.force _False) ff then let patts=patterns_of_constr env sigma nrels atom in `Nrule patts else - quantified_atom_of_constr (Environ.push_rel (RelDecl.LocalAssum (id,atom)) env) sigma (succ nrels) ff + quantified_atom_of_constr (Environ.push_rel (local_assum (id,atom)) env) sigma (succ nrels) ff | _ -> let patts=patterns_of_constr env sigma nrels term in `Rule patts let litteral_of_constr env sigma term= - match kind_of_term (whd_delta env term) with + match EConstr.kind sigma (whd_delta env sigma term) with | Prod (id,atom,ff) -> - if is_global (Lazy.force _False) ff then + if is_global sigma (Lazy.force _False) ff then match (atom_of_constr env sigma atom) with `Eq(t,a,b) -> `Neq(t,a,b) | `Other(p) -> `Nother(p) else begin try - quantified_atom_of_constr (Environ.push_rel (RelDecl.LocalAssum (id,atom)) env) sigma 1 ff + quantified_atom_of_constr (Environ.push_rel (local_assum (id,atom)) env) sigma 1 ff with Not_found -> `Other (decompose_term env sigma term) end @@ -197,8 +201,8 @@ let make_prb gls depth additionnal_terms = (fun decl -> let id = NamedDecl.get_id decl in begin - let cid=mkVar id in - match litteral_of_constr env sigma (NamedDecl.get_type decl) with + let cid=Constr.mkVar id in + match litteral_of_constr env sigma (EConstr.of_constr (NamedDecl.get_type decl)) with `Eq (t,a,b) -> add_equality state cid a b | `Neq (t,a,b) -> add_disequality state (Hyp cid) a b | `Other ph -> @@ -217,7 +221,7 @@ let make_prb gls depth additionnal_terms = | `Nrule patts -> add_quant state id false patts end) (Environ.named_context_of_val (Goal.V82.nf_hyps gls.sigma gls.it)); begin - match atom_of_constr env sigma (Evarutil.nf_evar sigma (pf_concl gls)) with + match atom_of_constr env sigma (EConstr.of_constr (pf_concl gls)) with `Eq (t,a,b) -> add_disequality state Goal a b | `Other g -> List.iter @@ -230,8 +234,7 @@ let make_prb gls depth additionnal_terms = let build_projection intype (cstr:pconstructor) special default gls= let ci= (snd(fst cstr)) in - let body=Equality.build_selector (pf_env gls) (project gls) ci (EConstr.mkRel 1) (EConstr.of_constr intype) (EConstr.of_constr special) default in - let body = EConstr.Unsafe.to_constr body in + let body=Equality.build_selector (pf_env gls) (project gls) ci (mkRel 1) intype special default in let id=pf_get_new_id (Id.of_string "t") gls in mkLambda(Name id,intype,body) @@ -240,10 +243,10 @@ let build_projection intype (cstr:pconstructor) special default gls= let _M =mkMeta let app_global f args k = - Tacticals.pf_constr_of_global (Lazy.force f) (fun fc -> k (EConstr.of_constr (mkApp (fc, args)))) + Tacticals.pf_constr_of_global (Lazy.force f) (fun fc -> k (mkApp (EConstr.of_constr fc, args))) let new_app_global f args k = - Tacticals.New.pf_constr_of_global (Lazy.force f) (fun fc -> k (EConstr.of_constr (mkApp (fc, args)))) + Tacticals.New.pf_constr_of_global (Lazy.force f) (fun fc -> k (mkApp (EConstr.of_constr fc, args))) let new_refine c = Proofview.V82.tactic (refine c) let refine c = refine c @@ -259,20 +262,24 @@ let refresh_type env evm ty = (Some false) env evm ty let refresh_universes ty k = - Proofview.Goal.enter { enter = begin fun gl -> + Proofview.Goal.s_enter { s_enter = begin fun gl -> let env = Proofview.Goal.env gl in let evm = Tacmach.New.project gl in - let evm, ty = refresh_type env evm (EConstr.of_constr ty) in - Tacticals.New.tclTHEN (Proofview.V82.tactic (Refiner.tclEVARS evm)) (k ty) + let evm, ty = refresh_type env evm ty in + let ty = EConstr.of_constr ty in + Sigma.Unsafe.of_pair (k ty, evm) end } +let constr_of_term c = EConstr.of_constr (constr_of_term c) + let rec proof_tac p : unit Proofview.tactic = Proofview.Goal.nf_enter { enter = begin fun gl -> - let type_of t = Tacmach.New.pf_unsafe_type_of gl (EConstr.of_constr t) in + let type_of t = EConstr.of_constr (Tacmach.New.pf_unsafe_type_of gl t) in try (* type_of can raise exceptions *) match p.p_rule with Ax c -> exact_check (EConstr.of_constr c) | SymAx c -> + let c = EConstr.of_constr c in let l=constr_of_term p.p_lhs and r=constr_of_term p.p_rhs in refresh_universes (type_of l) (fun typ -> @@ -321,7 +328,6 @@ let rec proof_tac p : unit Proofview.tactic = let special=mkRel (1+nargs-argind) in refresh_universes (type_of ti) (fun intype -> refresh_universes (type_of default) (fun outtype -> - let default = EConstr.of_constr default in let proj = Tacmach.New.of_old (build_projection intype cstr special default) gl in @@ -336,12 +342,11 @@ let refute_tac c t1 t2 p = let tt1=constr_of_term t1 and tt2=constr_of_term t2 in let hid = Tacmach.New.of_old (pf_get_new_id (Id.of_string "Heq")) gl in let false_t=mkApp (c,[|mkVar hid|]) in - let false_t = EConstr.of_constr false_t in let k intype = let neweq= new_app_global _eq [|intype;tt1;tt2|] in Tacticals.New.tclTHENS (neweq (assert_before (Name hid))) [proof_tac p; simplest_elim false_t] - in refresh_universes (Tacmach.New.pf_unsafe_type_of gl (EConstr.of_constr tt1)) k + in refresh_universes (EConstr.of_constr (Tacmach.New.pf_unsafe_type_of gl tt1)) k end } let refine_exact_check c gl = @@ -359,7 +364,7 @@ let convert_to_goal_tac c t1 t2 p = let endt=app_global _eq_rect [|sort;tt1;identity;c;tt2;mkVar e|] in Tacticals.New.tclTHENS (neweq (assert_before (Name e))) [proof_tac p; Proofview.V82.tactic (endt refine_exact_check)] - in refresh_universes (Tacmach.New.pf_unsafe_type_of gl (EConstr.of_constr tt2)) k + in refresh_universes (EConstr.of_constr (Tacmach.New.pf_unsafe_type_of gl tt2)) k end } let convert_to_hyp_tac c1 t1 c2 t2 p = @@ -367,8 +372,6 @@ let convert_to_hyp_tac c1 t1 c2 t2 p = let tt2=constr_of_term t2 in let h = Tacmach.New.of_old (pf_get_new_id (Id.of_string "H")) gl in let false_t=mkApp (c2,[|mkVar h|]) in - let false_t = EConstr.of_constr false_t in - let tt2 = EConstr.of_constr tt2 in Tacticals.New.tclTHENS (assert_before (Name h) tt2) [convert_to_goal_tac c1 t1 t2 p; simplest_elim false_t] @@ -381,20 +384,21 @@ let discriminate_tac (cstr,u as cstru) p = let concl = Proofview.Goal.concl gl in let xid = Tacmach.New.of_old (pf_get_new_id (Id.of_string "X")) gl in let identity = Universes.constr_of_global (Lazy.force _I) in + let identity = EConstr.of_constr identity in let trivial = Universes.constr_of_global (Lazy.force _True) in + let trivial = EConstr.of_constr trivial in let evm = Tacmach.New.project gl in - let evm, intype = refresh_type env evm (EConstr.of_constr (Tacmach.New.pf_unsafe_type_of gl (EConstr.of_constr t1))) in + let evm, intype = refresh_type env evm (EConstr.of_constr (Tacmach.New.pf_unsafe_type_of gl t1)) in + let intype = EConstr.of_constr intype in let evm, outtype = Evd.new_sort_variable Evd.univ_flexible evm in let outtype = mkSort outtype in let pred = mkLambda(Name xid,outtype,mkRel 1) in let hid = Tacmach.New.of_old (pf_get_new_id (Id.of_string "Heq")) gl in let proj = Tacmach.New.of_old (build_projection intype cstru trivial concl) gl in - let concl = EConstr.Unsafe.to_constr concl in let injt=app_global _f_equal [|intype;outtype;proj;t1;t2;mkVar hid|] in let endt k = injt (fun injt -> - let injt = EConstr.Unsafe.to_constr injt in app_global _eq_rect [|outtype;trivial;pred;identity;concl;injt|] k) in let neweq=new_app_global _eq [|intype;t1;t2|] in @@ -410,7 +414,7 @@ let build_term_to_complete uf meta pac = let real_args = List.map (fun i -> constr_of_term (term uf i)) pac.args in let dummy_args = List.rev (List.init pac.arity meta) in let all_args = List.rev_append real_args dummy_args in - applistc (mkConstructU cinfo.ci_constr) all_args + applist (mkConstructU cinfo.ci_constr, all_args) let cc_tactic depth additionnal_terms = Proofview.Goal.nf_enter { enter = begin fun gl -> @@ -448,7 +452,7 @@ let cc_tactic depth additionnal_terms = str "\"congruence with (" ++ prlist_with_sep (fun () -> str ")" ++ spc () ++ str "(") - (Termops.print_constr_env env) + (EConstr.Unsafe.to_constr %> Termops.print_constr_env env) terms_to_complete ++ str ")\"," end ++ @@ -459,10 +463,13 @@ let cc_tactic depth additionnal_terms = let ta=term uf dis.lhs and tb=term uf dis.rhs in match dis.rule with Goal -> proof_tac p - | Hyp id -> refute_tac id ta tb p + | Hyp id -> refute_tac (EConstr.of_constr id) ta tb p | HeqG id -> + let id = EConstr.of_constr id in convert_to_goal_tac id ta tb p | HeqnH (ida,idb) -> + let ida = EConstr.of_constr ida in + let idb = EConstr.of_constr idb in convert_to_hyp_tac ida ta idb tb p end } @@ -487,20 +494,23 @@ let congruence_tac depth l = let mk_eq f c1 c2 k = Tacticals.New.pf_constr_of_global (Lazy.force f) (fun fc -> + let fc = EConstr.of_constr fc in Proofview.Goal.enter { enter = begin fun gl -> let open Tacmach.New in - let evm, ty = pf_apply type_of gl (EConstr.of_constr c1) in - let evm, ty = Evarsolve.refresh_universes (Some false) (pf_env gl) evm (EConstr.of_constr ty) in + let evm, ty = pf_apply type_of gl c1 in + let ty = EConstr.of_constr ty in + let evm, ty = Evarsolve.refresh_universes (Some false) (pf_env gl) evm ty in + let ty = EConstr.of_constr ty in let term = mkApp (fc, [| ty; c1; c2 |]) in - let evm, _ = type_of (pf_env gl) evm (EConstr.of_constr term) in + let evm, _ = type_of (pf_env gl) evm term in Tacticals.New.tclTHEN (Proofview.V82.tactic (Refiner.tclEVARS evm)) - (k (EConstr.of_constr term)) + (k term) end }) let f_equal = Proofview.Goal.nf_enter { enter = begin fun gl -> let concl = Proofview.Goal.concl gl in - let concl = EConstr.Unsafe.to_constr concl in + let sigma = Tacmach.New.project gl in let cut_eq c1 c2 = try (* type_of can raise an exception *) Tacticals.New.tclTHENS @@ -509,9 +519,9 @@ let f_equal = with e when Proofview.V82.catchable_exception e -> Proofview.tclZERO e in Proofview.tclORELSE - begin match kind_of_term concl with - | App (r,[|_;t;t'|]) when Globnames.is_global (Lazy.force _eq) r -> - begin match kind_of_term t, kind_of_term t' with + begin match EConstr.kind sigma concl with + | App (r,[|_;t;t'|]) when is_global sigma (Lazy.force _eq) r -> + begin match EConstr.kind sigma t, EConstr.kind sigma t' with | App (f,v), App (f',v') when Int.equal (Array.length v) (Array.length v') -> let rec cuts i = if i < 0 then Tacticals.New.tclTRY (congruence_tac 1000 []) diff --git a/plugins/cc/cctac.mli b/plugins/cc/cctac.mli index 7c1d9f1c0..de6eb982e 100644 --- a/plugins/cc/cctac.mli +++ b/plugins/cc/cctac.mli @@ -7,7 +7,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -open Term +open EConstr open Proof_type val proof_tac: Ccproof.proof -> unit Proofview.tactic diff --git a/plugins/cc/g_congruence.ml4 b/plugins/cc/g_congruence.ml4 index 6f6811334..b787e824f 100644 --- a/plugins/cc/g_congruence.ml4 +++ b/plugins/cc/g_congruence.ml4 @@ -18,9 +18,9 @@ DECLARE PLUGIN "cc_plugin" TACTIC EXTEND cc [ "congruence" ] -> [ congruence_tac 1000 [] ] |[ "congruence" integer(n) ] -> [ congruence_tac n [] ] - |[ "congruence" "with" ne_constr_list(l) ] -> [ congruence_tac 1000 l ] + |[ "congruence" "with" ne_constr_list(l) ] -> [ congruence_tac 1000 (List.map EConstr.of_constr l) ] |[ "congruence" integer(n) "with" ne_constr_list(l) ] -> - [ congruence_tac n l ] + [ congruence_tac n (List.map EConstr.of_constr l) ] END TACTIC EXTEND f_equal diff --git a/pretyping/reductionops.ml b/pretyping/reductionops.ml index 45e7abcb7..c796a91ca 100644 --- a/pretyping/reductionops.ml +++ b/pretyping/reductionops.ml @@ -1221,6 +1221,14 @@ let clos_norm_flags flgs env sigma t = (CClosure.inject (EConstr.Unsafe.to_constr t))) with e when is_anomaly e -> error "Tried to normalize ill-typed term" +let clos_whd_flags flgs env sigma t = + try + let evars ev = safe_evar_value sigma ev in + EConstr.of_constr (CClosure.whd_val + (CClosure.create_clos_infos ~evars flgs env) + (CClosure.inject (EConstr.Unsafe.to_constr t))) + with e when is_anomaly e -> error "Tried to normalize ill-typed term" + let nf_beta = clos_norm_flags CClosure.beta (Global.env ()) let nf_betaiota = clos_norm_flags CClosure.betaiota (Global.env ()) let nf_betaiotazeta = clos_norm_flags CClosure.betaiotazeta (Global.env ()) diff --git a/pretyping/reductionops.mli b/pretyping/reductionops.mli index add1d186b..8aaeeb2c2 100644 --- a/pretyping/reductionops.mli +++ b/pretyping/reductionops.mli @@ -149,6 +149,7 @@ val iterate_whd_gen : bool -> CClosure.RedFlags.reds -> (** {6 Generic Optimized Reduction Function using Closures } *) val clos_norm_flags : CClosure.RedFlags.reds -> reduction_function +val clos_whd_flags : CClosure.RedFlags.reds -> reduction_function (** Same as [(strong whd_beta[delta][iota])], but much faster on big terms *) val nf_beta : local_reduction_function -- cgit v1.2.3 From ffb59901f568351401f2f3d1f3334031658b8880 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Thu, 24 Nov 2016 14:22:54 +0100 Subject: Setoid_ring API using EConstr. --- plugins/setoid_ring/g_newring.ml4 | 20 +++-- plugins/setoid_ring/newring.ml | 183 +++++++++++++++++++++----------------- plugins/setoid_ring/newring.mli | 41 ++------- 3 files changed, 117 insertions(+), 127 deletions(-) diff --git a/plugins/setoid_ring/g_newring.ml4 b/plugins/setoid_ring/g_newring.ml4 index 89ec4c1c5..13cf8330b 100644 --- a/plugins/setoid_ring/g_newring.ml4 +++ b/plugins/setoid_ring/g_newring.ml4 @@ -77,9 +77,7 @@ END VERNAC COMMAND EXTEND AddSetoidRing CLASSIFIED AS SIDEFF | [ "Add" "Ring" ident(id) ":" constr(t) ring_mods_opt(l) ] -> - [ let l = match l with None -> [] | Some l -> l in - let (k,set,cst,pre,post,power,sign, div) = process_ring_mods l in - add_theory id (ic t) set k cst (pre,post) power sign div] + [ let l = match l with None -> [] | Some l -> l in add_theory id t l] | [ "Print" "Rings" ] => [Vernac_classifier.classify_as_query] -> [ Feedback.msg_notice (strbrk "The following ring structures have been declared:"); Spmap.iter (fun fn fi -> @@ -92,7 +90,11 @@ END TACTIC EXTEND ring_lookup | [ "ring_lookup" tactic0(f) "[" constr_list(lH) "]" ne_constr_list(lrt) ] -> - [ let (t,lr) = List.sep_last lrt in ring_lookup f lH lr t] + [ + let lH = List.map EConstr.of_constr lH in + let lrt = List.map EConstr.of_constr lrt in + let (t,lr) = List.sep_last lrt in ring_lookup f lH lr t + ] END let pr_field_mod = function @@ -114,9 +116,7 @@ END VERNAC COMMAND EXTEND AddSetoidField CLASSIFIED AS SIDEFF | [ "Add" "Field" ident(id) ":" constr(t) field_mods_opt(l) ] -> - [ let l = match l with None -> [] | Some l -> l in - let (k,set,inj,cst_tac,pre,post,power,sign,div) = process_field_mods l in - add_field_theory id (ic t) set k cst_tac inj (pre,post) power sign div] + [ let l = match l with None -> [] | Some l -> l in add_field_theory id t l ] | [ "Print" "Fields" ] => [Vernac_classifier.classify_as_query] -> [ Feedback.msg_notice (strbrk "The following field structures have been declared:"); Spmap.iter (fun fn fi -> @@ -129,5 +129,9 @@ END TACTIC EXTEND field_lookup | [ "field_lookup" tactic(f) "[" constr_list(lH) "]" ne_constr_list(lt) ] -> - [ let (t,l) = List.sep_last lt in field_lookup f lH l t ] + [ + let lH = List.map EConstr.of_constr lH in + let lt = List.map EConstr.of_constr lt in + let (t,l) = List.sep_last lt in field_lookup f lH l t + ] END diff --git a/plugins/setoid_ring/newring.ml b/plugins/setoid_ring/newring.ml index 2f6c00c9d..63eccaa40 100644 --- a/plugins/setoid_ring/newring.ml +++ b/plugins/setoid_ring/newring.ml @@ -11,6 +11,7 @@ open CErrors open Util open Names open Term +open EConstr open Vars open CClosure open Environ @@ -42,9 +43,9 @@ let tag_arg tag_rec map subs i c = | Prot -> mk_atom c | Rec -> if Int.equal i (-1) then mk_clos subs c else tag_rec c -let global_head_of_constr c = - let f, args = decompose_app c in - try global_of_constr f +let global_head_of_constr sigma c = + let f, args = decompose_app sigma c in + try fst (Termops.global_of_constr sigma f) with Not_found -> anomaly (str "global_head_of_constr") let global_of_constr_nofail c = @@ -52,6 +53,7 @@ let global_of_constr_nofail c = with Not_found -> VarRef (Id.of_string "dummy") let rec mk_clos_but f_map subs t = + let open Term in match f_map (global_of_constr_nofail t) with | Some map -> tag_arg (mk_clos_but f_map subs) map subs (-1) t | None -> @@ -61,6 +63,7 @@ let rec mk_clos_but f_map subs t = | _ -> mk_atom t) and mk_clos_app_but f_map subs f args n = + let open Term in if n >= Array.length args then mk_atom(mkApp(f, args)) else let fargs, args' = Array.chop n args in @@ -81,10 +84,11 @@ let lookup_map map = with Not_found -> user_err ~hdr:"lookup_map" (str"map "++qs map++str"not found") -let protect_red map env sigma c = - let c = EConstr.Unsafe.to_constr c in - EConstr.of_constr (kl (create_clos_infos all env) - (mk_clos_but (lookup_map map c) (Esubst.subs_id 0) c));; +let protect_red map env sigma c0 = + let evars ev = Evarutil.safe_evar_value sigma ev in + let c = EConstr.Unsafe.to_constr c0 in + EConstr.of_constr (kl (create_clos_infos ~evars all env) + (mk_clos_but (lookup_map map sigma c0) (Esubst.subs_id 0) c));; let protect_tac map = Tactics.reduct_option (protect_red map,DEFAULTcast) None @@ -137,14 +141,16 @@ let _ = add_tacdef false ((Loc.ghost,Id.of_string"ring_closed_term" let ic c = let env = Global.env() in let sigma = Evd.from_env env in - Constrintern.interp_open_constr env sigma c + let sigma, c = Constrintern.interp_open_constr env sigma c in + (sigma, EConstr.of_constr c) let ic_unsafe c = (*FIXME remove *) let env = Global.env() in let sigma = Evd.from_env env in - fst (Constrintern.interp_constr env sigma c) + EConstr.of_constr (fst (Constrintern.interp_constr env sigma c)) let decl_constant na ctx c = + let open Constr in let vars = Universes.universes_of_constr c in let ctx = Universes.restrict_universe_context (Univ.ContextSet.of_context ctx) vars in mkConst(declare_constant (Id.of_string na) @@ -197,6 +203,7 @@ let get_res = let exec_tactic env evd n f args = let fold arg (i, vars, lfun) = + let arg = EConstr.Unsafe.to_constr arg in let id = Id.of_string ("x" ^ string_of_int i) in let x = Reference (ArgVar (Loc.ghost, id)) in (succ i, x :: vars, Id.Map.add id (Value.of_constr arg) lfun) @@ -212,7 +219,8 @@ let exec_tactic env evd n f args = let gl = dummy_goal env evd in let gls = Proofview.V82.of_tactic (Tacinterp.eval_tactic_ist ist (ltac_call f (args@[getter]))) gl in let evd, nf = Evarutil.nf_evars_and_universes (Refiner.project gls) in - Array.map (fun x -> nf (constr_of x)) !tactic_res, snd (Evd.universe_context evd) + let nf c = nf (constr_of c) in + Array.map nf !tactic_res, snd (Evd.universe_context evd) let stdlib_modules = [["Coq";"Setoids";"Setoid"]; @@ -222,7 +230,7 @@ let stdlib_modules = ] let coq_constant c = - lazy (Coqlib.gen_constant_in_modules "Ring" stdlib_modules c) + lazy (EConstr.of_constr (Coqlib.gen_constant_in_modules "Ring" stdlib_modules c)) let coq_reference c = lazy (Coqlib.gen_reference_in_modules "Ring" stdlib_modules c) @@ -238,21 +246,22 @@ let lapp f args = mkApp(Lazy.force f,args) let plapp evd f args = let fc = Evarutil.e_new_global evd (Lazy.force f) in + let fc = EConstr.of_constr fc in mkApp(fc,args) -let dest_rel0 t = - match kind_of_term t with +let dest_rel0 sigma t = + match EConstr.kind sigma t with | App(f,args) when Array.length args >= 2 -> let rel = mkApp(f,Array.sub args 0 (Array.length args - 2)) in - if closed0 rel then + if closed0 sigma rel then (rel,args.(Array.length args - 2),args.(Array.length args - 1)) else error "ring: cannot find relation (not closed)" | _ -> error "ring: cannot find relation" -let rec dest_rel t = - match kind_of_term t with - | Prod(_,_,c) -> dest_rel c - | _ -> dest_rel0 t +let rec dest_rel sigma t = + match EConstr.kind sigma t with + | Prod(_,_,c) -> dest_rel sigma c + | _ -> dest_rel0 sigma t (****************************************************************************) (* Library linking *) @@ -267,7 +276,7 @@ let plugin_modules = ] let my_constant c = - lazy (Coqlib.gen_constant_in_modules "Ring" plugin_modules c) + lazy (EConstr.of_constr (Coqlib.gen_constant_in_modules "Ring" plugin_modules c)) let my_reference c = lazy (Coqlib.gen_reference_in_modules "Ring" plugin_modules c) @@ -311,13 +320,13 @@ let coq_mkhypo = my_reference "mkhypo" let coq_hypo = my_reference "hypo" (* Equality: do not evaluate but make recursive call on both sides *) -let map_with_eq arg_map c = - let (req,_,_) = dest_rel c in +let map_with_eq arg_map sigma c = + let (req,_,_) = dest_rel sigma c in interp_map - ((global_head_of_constr req,(function -1->Prot|_->Rec)):: + ((global_head_of_constr sigma req,(function -1->Prot|_->Rec)):: List.map (fun (c,map) -> (Lazy.force c,map)) arg_map) -let map_without_eq arg_map _ = +let map_without_eq arg_map _ _ = interp_map (List.map (fun (c,map) -> (Lazy.force c,map)) arg_map) let _ = add_map "ring" @@ -346,9 +355,9 @@ let ring_for_carrier r = Cmap.find r !from_carrier let find_ring_structure env sigma l = match l with | t::cl' -> - let ty = Retyping.get_type_of env sigma (EConstr.of_constr t) in + let ty = Retyping.get_type_of env sigma t in let check c = - let ty' = Retyping.get_type_of env sigma (EConstr.of_constr c) in + let ty' = Retyping.get_type_of env sigma c in if not (Reductionops.is_conv env sigma (EConstr.of_constr ty) (EConstr.of_constr ty')) then user_err ~hdr:"ring" (str"arguments of ring_simplify do not have all the same type") @@ -381,7 +390,7 @@ let subst_th (subst,th) = let posttac'= Tacsubst.subst_tactic subst th.ring_post_tac in if c' == th.ring_carrier && eq' == th.ring_req && - eq_constr set' th.ring_setoid && + Term.eq_constr set' th.ring_setoid && ext' == th.ring_ext && morph' == th.ring_morph && th' == th.ring_th && @@ -417,18 +426,11 @@ let theory_to_obj : ring_info -> obj = let setoid_of_relation env evd a r = - let a = EConstr.of_constr a in - let r = EConstr.of_constr r in try let evm = !evd in let evm, refl = Rewrite.get_reflexive_proof env evm a r in let evm, sym = Rewrite.get_symmetric_proof env evm a r in let evm, trans = Rewrite.get_transitive_proof env evm a r in - let refl = EConstr.Unsafe.to_constr refl in - let sym = EConstr.Unsafe.to_constr sym in - let trans = EConstr.Unsafe.to_constr trans in - let a = EConstr.Unsafe.to_constr a in - let r = EConstr.Unsafe.to_constr r in evd := evm; lapp coq_mk_Setoid [|a ; r ; refl; sym; trans |] with Not_found -> @@ -494,44 +496,37 @@ let op_smorph r add mul req m1 m2 = (* (setoid,op_morph) *) let ring_equality env evd (r,add,mul,opp,req) = - match kind_of_term req with - | App (f, [| _ |]) when eq_constr_nounivs f (Lazy.force coq_eq) -> + let pr_constr c = pr_constr (EConstr.to_constr !evd c) in + match EConstr.kind !evd req with + | App (f, [| _ |]) when eq_constr_nounivs !evd f (Lazy.force coq_eq) -> let setoid = plapp evd coq_eq_setoid [|r|] in let op_morph = match opp with Some opp -> plapp evd coq_eq_morph [|r;add;mul;opp|] | None -> plapp evd coq_eq_smorph [|r;add;mul|] in - let setoid = Typing.e_solve_evars env evd (EConstr.of_constr setoid) in - let op_morph = Typing.e_solve_evars env evd (EConstr.of_constr op_morph) in + let setoid = Typing.e_solve_evars env evd setoid in + let op_morph = Typing.e_solve_evars env evd op_morph in + let setoid = EConstr.of_constr setoid in + let op_morph = EConstr.of_constr op_morph in (setoid,op_morph) | _ -> let setoid = setoid_of_relation (Global.env ()) evd r req in - let signature = - let r = EConstr.of_constr r in - let req = EConstr.of_constr req in - [Some (r,Some req);Some (r,Some req)],Some(r,Some req) - in -(* let signature = [Some (r,Some req);Some (r,Some req)],Some(r,Some req) in *) + let signature = [Some (r,Some req);Some (r,Some req)],Some(r,Some req) in let add_m, add_m_lem = - try Rewrite.default_morphism signature (EConstr.of_constr add) + try Rewrite.default_morphism signature add with Not_found -> error "ring addition should be declared as a morphism" in - let add_m_lem = EConstr.Unsafe.to_constr add_m_lem in let mul_m, mul_m_lem = - try Rewrite.default_morphism signature (EConstr.of_constr mul) + try Rewrite.default_morphism signature mul with Not_found -> error "ring multiplication should be declared as a morphism" in - let mul_m_lem = EConstr.Unsafe.to_constr mul_m_lem in let op_morph = match opp with | Some opp -> (let opp_m,opp_m_lem = - let r = EConstr.of_constr r in - let req = EConstr.of_constr req in - try Rewrite.default_morphism ([Some(r,Some req)],Some(r,Some req)) (EConstr.of_constr opp) + try Rewrite.default_morphism ([Some(r,Some req)],Some(r,Some req)) opp with Not_found -> error "ring opposite should be declared as a morphism" in - let opp_m_lem = EConstr.Unsafe.to_constr opp_m_lem in let op_morph = op_morph r add mul opp req add_m_lem mul_m_lem opp_m_lem in Flags.if_verbose @@ -558,16 +553,17 @@ let build_setoid_params env evd r add mul opp req eqth = | None -> ring_equality env evd (r,add,mul,opp,req) let dest_ring env sigma th_spec = - let th_typ = Retyping.get_type_of env sigma (EConstr.of_constr th_spec) in - match kind_of_term th_typ with + let th_typ = Retyping.get_type_of env sigma th_spec in + let th_typ = EConstr.of_constr th_typ in + match EConstr.kind sigma th_typ with App(f,[|r;zero;one;add;mul;sub;opp;req|]) - when eq_constr_nounivs f (Lazy.force coq_almost_ring_theory) -> + when eq_constr_nounivs sigma f (Lazy.force coq_almost_ring_theory) -> (None,r,zero,one,add,mul,Some sub,Some opp,req) | App(f,[|r;zero;one;add;mul;req|]) - when eq_constr_nounivs f (Lazy.force coq_semi_ring_theory) -> + when eq_constr_nounivs sigma f (Lazy.force coq_semi_ring_theory) -> (Some true,r,zero,one,add,mul,None,None,req) | App(f,[|r;zero;one;add;mul;sub;opp;req|]) - when eq_constr_nounivs f (Lazy.force coq_ring_theory) -> + when eq_constr_nounivs sigma f (Lazy.force coq_ring_theory) -> (Some false,r,zero,one,add,mul,Some sub,Some opp,req) | _ -> error "bad ring structure" @@ -589,21 +585,24 @@ let interp_cst_tac env sigma rk kind (zero,one,add,mul,opp) cst_tac = TacArg(Loc.ghost,TacCall(Loc.ghost,t,[])) let make_hyp env evd c = - let t = Retyping.get_type_of env !evd (EConstr.of_constr c) in + let t = Retyping.get_type_of env !evd c in + let t = EConstr.of_constr t in plapp evd coq_mkhypo [|t;c|] let make_hyp_list env evd lH = let carrier = Evarutil.e_new_global evd (Lazy.force coq_hypo) in + let carrier = EConstr.of_constr carrier in let l = List.fold_right (fun c l -> plapp evd coq_cons [|carrier; (make_hyp env evd c); l|]) lH (plapp evd coq_nil [|carrier|]) in - let l' = Typing.e_solve_evars env evd (EConstr.of_constr l) in + let l' = Typing.e_solve_evars env evd l in Evarutil.nf_evars_universes !evd l' let interp_power env evd pow = let carrier = Evarutil.e_new_global evd (Lazy.force coq_hypo) in + let carrier = EConstr.of_constr carrier in match pow with | None -> let t = ArgArg(Loc.ghost, Lazy.force ltac_inv_morph_nothing) in @@ -619,6 +618,7 @@ let interp_power env evd pow = let interp_sign env evd sign = let carrier = Evarutil.e_new_global evd (Lazy.force coq_hypo) in + let carrier = EConstr.of_constr carrier in match sign with | None -> plapp evd coq_None [|carrier|] | Some spec -> @@ -628,6 +628,7 @@ let interp_sign env evd sign = let interp_div env evd div = let carrier = Evarutil.e_new_global evd (Lazy.force coq_hypo) in + let carrier = EConstr.of_constr carrier in match div with | None -> plapp evd coq_None [|carrier|] | Some spec -> @@ -635,7 +636,7 @@ let interp_div env evd div = plapp evd coq_Some [|carrier;spec|] (* Same remark on ill-typed terms ... *) -let add_theory name (sigma,rth) eqth morphth cst_tac (pre,post) power sign div = +let add_theory0 name (sigma, rth) eqth morphth cst_tac (pre,post) power sign div = check_required_library (cdir@["Ring_base"]); let env = Global.env() in let (kind,r,zero,one,add,mul,sub,opp,req) = dest_ring env sigma rth in @@ -665,6 +666,9 @@ let add_theory name (sigma,rth) eqth morphth cst_tac (pre,post) power sign div = match post with Some t -> Tacintern.glob_tactic t | _ -> TacId [] in + let r = EConstr.to_constr sigma r in + let req = EConstr.to_constr sigma req in + let sth = EConstr.to_constr sigma sth in let _ = Lib.add_leaf name (theory_to_obj @@ -712,20 +716,25 @@ let process_ring_mods l = let k = match !kind with Some k -> k | None -> Abstract in (k, !set, !cst_tac, !pre, !post, !power, !sign, !div) +let add_theory id rth l = + let (sigma, rth) = ic rth in + let (k,set,cst,pre,post,power,sign, div) = process_ring_mods l in + add_theory0 id (sigma, rth) set k cst (pre,post) power sign div + (*****************************************************************************) (* The tactics consist then only in a lookup in the ring database and call the appropriate ltac. *) -let make_args_list rl t = +let make_args_list sigma rl t = match rl with - | [] -> let (_,t1,t2) = dest_rel0 t in [t1;t2] + | [] -> let (_,t1,t2) = dest_rel0 sigma t in [t1;t2] | _ -> rl let make_term_list env evd carrier rl = let l = List.fold_right (fun x l -> plapp evd coq_cons [|carrier;x;l|]) rl (plapp evd coq_nil [|carrier|]) - in Typing.e_solve_evars env evd (EConstr.of_constr l) + in Typing.e_solve_evars env evd l let carg = Tacinterp.Value.of_constr let tacarg expr = @@ -751,10 +760,10 @@ let ring_lookup (f : Value.t) lH rl t = let sigma = Tacmach.New.project gl in let env = Proofview.Goal.env gl in try (* find_ring_strucure can raise an exception *) + let rl = make_args_list sigma rl t in let evdref = ref sigma in - let rl = make_args_list rl t in let e = find_ring_structure env sigma rl in - let rl = carg (make_term_list env evdref e.ring_carrier rl) in + let rl = carg (make_term_list env evdref (EConstr.of_constr e.ring_carrier) rl) in let lH = carg (make_hyp_list env evdref lH) in let ring = ltac_ring_structure e in Proofview.tclTHEN (Proofview.Unsafe.tclEVARS !evdref) (ltac_apply f (ring@[lH;rl])) @@ -814,21 +823,23 @@ let af_ar = my_reference"AF_AR" let f_r = my_reference"F_R" let sf_sr = my_reference"SF_SR" let dest_field env evd th_spec = - let th_typ = Retyping.get_type_of env !evd (EConstr.of_constr th_spec) in - match kind_of_term th_typ with + let open Termops in + let th_typ = Retyping.get_type_of env !evd th_spec in + let th_typ = EConstr.of_constr th_typ in + match EConstr.kind !evd th_typ with | App(f,[|r;zero;one;add;mul;sub;opp;div;inv;req|]) - when is_global (Lazy.force afield_theory) f -> + when is_global !evd (Lazy.force afield_theory) f -> let rth = plapp evd af_ar [|r;zero;one;add;mul;sub;opp;div;inv;req;th_spec|] in (None,r,zero,one,add,mul,Some sub,Some opp,div,inv,req,rth) | App(f,[|r;zero;one;add;mul;sub;opp;div;inv;req|]) - when is_global (Lazy.force field_theory) f -> + when is_global !evd (Lazy.force field_theory) f -> let rth = plapp evd f_r [|r;zero;one;add;mul;sub;opp;div;inv;req;th_spec|] in (Some false,r,zero,one,add,mul,Some sub,Some opp,div,inv,req,rth) | App(f,[|r;zero;one;add;mul;div;inv;req|]) - when is_global (Lazy.force sfield_theory) f -> + when is_global !evd (Lazy.force sfield_theory) f -> let rth = plapp evd sf_sr [|r;zero;one;add;mul;div;inv;req;th_spec|] in (Some true,r,zero,one,add,mul,None,None,div,inv,req,rth) @@ -843,9 +854,9 @@ let find_field_structure env sigma l = check_required_library (cdir@["Field_tac"]); match l with | t::cl' -> - let ty = Retyping.get_type_of env sigma (EConstr.of_constr t) in + let ty = Retyping.get_type_of env sigma t in let check c = - let ty' = Retyping.get_type_of env sigma (EConstr.of_constr c) in + let ty' = Retyping.get_type_of env sigma c in if not (Reductionops.is_conv env sigma (EConstr.of_constr ty) (EConstr.of_constr ty')) then user_err ~hdr:"field" (str"arguments of field_simplify do not have all the same type") @@ -908,31 +919,31 @@ let ftheory_to_obj : field_info -> obj = classify_function = (fun x -> Substitute x) } let field_equality evd r inv req = - match kind_of_term req with - | App (f, [| _ |]) when eq_constr_nounivs f (Lazy.force coq_eq) -> - mkApp(Universes.constr_of_global (Coqlib.build_coq_eq_data()).congr,[|r;r;inv|]) + match EConstr.kind !evd req with + | App (f, [| _ |]) when eq_constr_nounivs !evd f (Lazy.force coq_eq) -> + let c = Universes.constr_of_global (Coqlib.build_coq_eq_data()).congr in + let c = EConstr.of_constr c in + mkApp(c,[|r;r;inv|]) | _ -> let _setoid = setoid_of_relation (Global.env ()) evd r req in - let r = EConstr.of_constr r in - let req = EConstr.of_constr req in let signature = [Some (r,Some req)],Some(r,Some req) in - let inv = EConstr.of_constr inv in let inv_m, inv_m_lem = try Rewrite.default_morphism signature inv with Not_found -> error "field inverse should be declared as a morphism" in - let inv_m_lem = EConstr.Unsafe.to_constr inv_m_lem in inv_m_lem -let add_field_theory name (sigma,fth) eqth morphth cst_tac inj (pre,post) power sign odiv = +let add_field_theory0 name fth eqth morphth cst_tac inj (pre,post) power sign odiv = + let open Constr in check_required_library (cdir@["Field_tac"]); + let (sigma,fth) = ic fth in let env = Global.env() in let evd = ref sigma in let (kind,r,zero,one,add,mul,sub,opp,div,inv,req,rth) = dest_field env evd fth in let (sth,ext) = build_setoid_params env evd r add mul opp req eqth in let eqth = Some(sth,ext) in - let _ = add_theory name (!evd,rth) eqth morphth cst_tac (None,None) power sign odiv in + let _ = add_theory0 name (!evd,rth) eqth morphth cst_tac (None,None) power sign odiv in let (pow_tac, pspec) = interp_power env evd power in let sspec = interp_sign env evd sign in let dspec = interp_div env evd odiv in @@ -947,7 +958,7 @@ let add_field_theory name (sigma,fth) eqth morphth cst_tac inj (pre,post) power let lemma4 = params.(6) in let cond_lemma = match inj with - | Some thm -> mkApp(params.(8),[|thm|]) + | Some thm -> mkApp(params.(8),[|EConstr.to_constr sigma thm|]) | None -> params.(7) in let lemma1 = decl_constant (Id.to_string name^"_field_lemma1") ctx lemma1 in @@ -969,6 +980,8 @@ let add_field_theory name (sigma,fth) eqth morphth cst_tac inj (pre,post) power match post with Some t -> Tacintern.glob_tactic t | _ -> TacId [] in + let r = EConstr.to_constr sigma r in + let req = EConstr.to_constr sigma req in let _ = Lib.add_leaf name (ftheory_to_obj @@ -1008,6 +1021,10 @@ let process_field_mods l = let k = match !kind with Some k -> k | None -> Abstract in (k, !set, !inj, !cst_tac, !pre, !post, !power, !sign, !div) +let add_field_theory id t mods = + let (k,set,inj,cst_tac,pre,post,power,sign,div) = process_field_mods mods in + add_field_theory0 id t set k cst_tac inj (pre,post) power sign div + let ltac_field_structure e = let req = carg e.field_req in let cst_tac = tacarg e.field_cst_tac in @@ -1027,10 +1044,10 @@ let field_lookup (f : Value.t) lH rl t = let sigma = Tacmach.New.project gl in let env = Proofview.Goal.env gl in try + let rl = make_args_list sigma rl t in let evdref = ref sigma in - let rl = make_args_list rl t in let e = find_field_structure env sigma rl in - let rl = carg (make_term_list env evdref e.field_carrier rl) in + let rl = carg (make_term_list env evdref (EConstr.of_constr e.field_carrier) rl) in let lH = carg (make_hyp_list env evdref lH) in let field = ltac_field_structure e in Proofview.tclTHEN (Proofview.Unsafe.tclEVARS !evdref) (ltac_apply f (field@[lH;rl])) diff --git a/plugins/setoid_ring/newring.mli b/plugins/setoid_ring/newring.mli index 89538eb24..4367d021c 100644 --- a/plugins/setoid_ring/newring.mli +++ b/plugins/setoid_ring/newring.mli @@ -8,6 +8,7 @@ open Names open Constr +open EConstr open Libnames open Globnames open Constrexpr @@ -21,26 +22,10 @@ val protect_tac : string -> unit Proofview.tactic val closed_term : EConstr.constr -> global_reference list -> unit Proofview.tactic -val process_ring_mods : - constr_expr ring_mod list -> - constr coeff_spec * (constr * constr) option * - cst_tac_spec option * raw_tactic_expr option * - raw_tactic_expr option * - (cst_tac_spec * constr_expr) option * - constr_expr option * constr_expr option - val add_theory : Id.t -> - Evd.evar_map * constr -> - (constr * constr) option -> - constr coeff_spec -> - cst_tac_spec option -> - raw_tactic_expr option * raw_tactic_expr option -> - (cst_tac_spec * constr_expr) option -> - constr_expr option -> - constr_expr option -> unit - -val ic : constr_expr -> Evd.evar_map * constr + constr_expr -> + constr_expr ring_mod list -> unit val from_name : ring_info Spmap.t ref @@ -49,26 +34,10 @@ val ring_lookup : constr list -> constr list -> constr -> unit Proofview.tactic -val process_field_mods : - constr_expr field_mod list -> - constr coeff_spec * - (constr * constr) option * constr option * - cst_tac_spec option * raw_tactic_expr option * - raw_tactic_expr option * - (cst_tac_spec * constr_expr) option * - constr_expr option * constr_expr option - val add_field_theory : Id.t -> - Evd.evar_map * constr -> - (constr * constr) option -> - constr coeff_spec -> - cst_tac_spec option -> - constr option -> - raw_tactic_expr option * raw_tactic_expr option -> - (cst_tac_spec * constr_expr) option -> - constr_expr option -> - constr_expr option -> unit + constr_expr -> + constr_expr field_mod list -> unit val field_from_name : field_info Spmap.t ref -- cgit v1.2.3 From b36adb2124d3ba8a5547605e7f89bb0835d0ab10 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Thu, 24 Nov 2016 15:50:17 +0100 Subject: Removing some return type compatibility layers in Termops. --- engine/evarutil.ml | 11 +- engine/evarutil.mli | 16 +-- engine/termops.ml | 70 +++++----- engine/termops.mli | 185 ++++++++++++------------- ltac/rewrite.ml | 3 - plugins/cc/cctac.ml | 2 - plugins/decl_mode/decl_interp.ml | 2 +- plugins/firstorder/formula.ml | 4 +- plugins/firstorder/unify.ml | 4 +- plugins/funind/functional_principles_proofs.ml | 29 ++-- plugins/funind/functional_principles_types.ml | 20 +-- plugins/funind/invfun.ml | 8 +- plugins/funind/merge.ml | 7 +- plugins/funind/recdef.ml | 3 +- plugins/quote/quote.ml | 2 +- plugins/setoid_ring/newring.ml | 5 - pretyping/cases.ml | 1 - pretyping/detyping.ml | 8 +- pretyping/evarconv.ml | 2 +- pretyping/evarsolve.ml | 4 +- pretyping/inductiveops.ml | 4 +- pretyping/program.ml | 2 +- pretyping/reductionops.ml | 8 +- printing/prettyp.ml | 2 +- stm/lemmas.ml | 2 +- tactics/autorewrite.ml | 2 +- tactics/class_tactics.ml | 2 +- tactics/eqschemes.ml | 2 +- tactics/hints.ml | 4 +- tactics/tactics.ml | 2 +- tactics/term_dnet.ml | 4 +- toplevel/command.ml | 18 +-- toplevel/obligations.ml | 3 +- toplevel/record.ml | 2 +- 34 files changed, 223 insertions(+), 220 deletions(-) diff --git a/engine/evarutil.ml b/engine/evarutil.ml index 73286f2c4..8940be09d 100644 --- a/engine/evarutil.ml +++ b/engine/evarutil.ml @@ -47,10 +47,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 Sigma (c, sigma, p) = Sigma.fresh_global (Global.env()) evd x in + Sigma (EConstr.of_constr c, sigma, p) (****************************************************) (* Expanding/testing/exposing existential variables *) @@ -230,7 +231,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) *) @@ -396,6 +397,7 @@ let default_source = (Loc.ghost,Evar_kinds.InternalHole) let restrict_evar evd evk filter candidates = let evd = Sigma.to_evar_map evd in + let candidates = Option.map (fun l -> List.map EConstr.Unsafe.to_constr l) candidates in let evd, evk' = Evd.restrict evk filter ?candidates evd in Sigma.Unsafe.of_pair (evk', Evd.declare_future_goal evk' evd) @@ -408,6 +410,7 @@ let new_pure_evar_full evd evi = let new_pure_evar sign evd ?(src=default_source) ?(filter = Filter.identity) ?candidates ?(store = Store.empty) ?naming ?(principal=false) typ = let typ = EConstr.Unsafe.to_constr typ in let evd = Sigma.to_evar_map evd 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 evi = { @@ -437,7 +440,7 @@ let new_evar_instance sign evd typ ?src ?filter ?candidates ?store ?naming ?prin (* 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 map c = EConstr.Unsafe.to_constr (subst2 subst vsubst (EConstr.of_constr c)) 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 diff --git a/engine/evarutil.mli b/engine/evarutil.mli index 5882f02b9..ea9599c8b 100644 --- a/engine/evarutil.mli +++ b/engine/evarutil.mli @@ -17,18 +17,18 @@ open Environ (** [new_meta] is a generator of unique meta variables *) val new_meta : unit -> metavariable -val mk_new_meta : unit -> constr +val mk_new_meta : unit -> EConstr.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 -> - ?candidates:constr list -> ?store:Store.t -> + ?candidates:EConstr.constr list -> ?store:Store.t -> ?naming:Misctypes.intro_pattern_naming_expr -> ?principal:bool -> EConstr.types -> (EConstr.constr, 'r) Sigma.sigma val new_pure_evar : named_context_val -> 'r Sigma.t -> ?src:Loc.t * Evar_kinds.t -> ?filter:Filter.t -> - ?candidates:constr list -> ?store:Store.t -> + ?candidates:EConstr.constr list -> ?store:Store.t -> ?naming:Misctypes.intro_pattern_naming_expr -> ?principal:bool -> EConstr.types -> (evar, 'r) Sigma.sigma @@ -37,7 +37,7 @@ val new_pure_evar_full : 'r Sigma.t -> evar_info -> (evar, 'r) Sigma.sigma (** the same with side-effects *) val e_new_evar : env -> evar_map ref -> ?src:Loc.t * Evar_kinds.t -> ?filter:Filter.t -> - ?candidates:constr list -> ?store:Store.t -> + ?candidates:EConstr.constr list -> ?store:Store.t -> ?naming:Misctypes.intro_pattern_naming_expr -> ?principal:bool -> EConstr.types -> EConstr.constr @@ -56,12 +56,12 @@ val new_Type : ?rigid:rigid -> env -> 'r Sigma.t -> (EConstr.constr, 'r) Sigma.s val e_new_Type : ?rigid:rigid -> env -> evar_map ref -> EConstr.constr val restrict_evar : 'r Sigma.t -> existential_key -> Filter.t -> - constr list option -> (existential_key, 'r) Sigma.sigma + EConstr.constr list option -> (existential_key, 'r) Sigma.sigma (** Polymorphic constants *) -val new_global : 'r Sigma.t -> Globnames.global_reference -> (constr, 'r) Sigma.sigma -val e_new_global : evar_map ref -> Globnames.global_reference -> constr +val new_global : 'r Sigma.t -> Globnames.global_reference -> (EConstr.constr, 'r) Sigma.sigma +val e_new_global : evar_map ref -> Globnames.global_reference -> EConstr.constr (** Create a fresh evar in a context different from its definition context: [new_evar_instance sign evd ty inst] creates a new evar of context @@ -71,7 +71,7 @@ val e_new_global : evar_map ref -> Globnames.global_reference -> constr as a telescope) is [sign] *) val new_evar_instance : named_context_val -> 'r Sigma.t -> EConstr.types -> - ?src:Loc.t * Evar_kinds.t -> ?filter:Filter.t -> ?candidates:constr list -> + ?src:Loc.t * Evar_kinds.t -> ?filter:Filter.t -> ?candidates:EConstr.constr list -> ?store:Store.t -> ?naming:Misctypes.intro_pattern_naming_expr -> ?principal:bool -> EConstr.constr list -> (EConstr.constr, 'r) Sigma.sigma diff --git a/engine/termops.ml b/engine/termops.ml index 879d77de2..465832c10 100644 --- a/engine/termops.ml +++ b/engine/termops.ml @@ -210,12 +210,7 @@ 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 @@ -628,9 +623,10 @@ let vars_of_global_reference env gr = [m] is appropriately lifted through abstractions of [t] *) let dependent_main noevar univs sigma m t = + let open EConstr in let eqc x y = - if univs then not (Option.is_empty (EConstr.eq_constr_universes sigma x y)) - else EConstr.eq_constr_nounivs sigma 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 @@ -638,13 +634,13 @@ let dependent_main noevar univs sigma m t = else 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 (EConstr.mkApp (ft,Array.sub lt 0 (Array.length lm))); + 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 && EConstr.isMeta sigma c -> () + | _, Cast (c,_,_) when noevar && isMeta sigma c -> () | _, Evar _ when noevar -> () - | _ -> EConstr.iter_with_binders sigma (fun c -> EConstr.Vars.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 @@ -661,6 +657,7 @@ let dependent_in_decl sigma a decl = | LocalDef (_, body, t) -> dependent sigma a (EConstr.of_constr body) || dependent sigma a (EConstr.of_constr t) let count_occurrences sigma m t = + let open EConstr in let n = ref 0 in let rec countrec m t = if EConstr.eq_constr sigma m t then @@ -668,13 +665,13 @@ let count_occurrences sigma m t = else 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 (EConstr.mkApp (ft,Array.sub lt 0 (Array.length lm))); + 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 EConstr.isMeta sigma c -> () + | _, Cast (c,_,_) when isMeta sigma c -> () | _, Evar _ -> () - | _ -> EConstr.iter_with_binders sigma (EConstr.Vars.lift 1) countrec m t + | _ -> EConstr.iter_with_binders sigma (Vars.lift 1) countrec m t in countrec m t; !n @@ -682,7 +679,7 @@ let count_occurrences sigma m t = (* Synonymous *) let occur_term = dependent -let pop t = EConstr.Unsafe.to_constr (EConstr.Vars.lift (-1) t) +let pop t = EConstr.Vars.lift (-1) t (***************************) (* bindings functions *) @@ -712,33 +709,35 @@ let collapse_appl sigma c = match EConstr.kind sigma c with | App (g,cl1) -> collapse_rec g (Array.append cl1 cl2) | _ -> EConstr.mkApp (f,cl2) in - EConstr.Unsafe.to_constr (collapse_rec f cl) - | _ -> EConstr.Unsafe.to_constr c + collapse_rec f cl + | _ -> c (* First utilities for avoiding telescope computation for subst_term *) let prefix_application sigma eq_fun (k,c) t = - let c' = EConstr.of_constr (collapse_appl sigma c) and t' = EConstr.of_constr (collapse_appl sigma t) in + 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 sigma c' (EConstr.mkApp (f2, Array.sub cl2 0 l1)) then - Some (EConstr.mkApp (EConstr.mkRel k, Array.sub cl2 l1 (l2 - l1))) + && 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 sigma eq_fun (k,c) by_c t = - let c' = EConstr.of_constr (collapse_appl sigma c) and t' = EConstr.of_constr (collapse_appl sigma t) in + 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 sigma c' (EConstr.mkApp (f2, Array.sub cl2 0 l1)) then - Some (EConstr.mkApp ((EConstr.Vars.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 @@ -754,7 +753,7 @@ let subst_term_gen sigma eq_fun c t = match prefix_application sigma eq_fun kc t with | Some x -> x | None -> - if eq_fun sigma c t then EConstr.mkRel k + if eq_fun sigma c t then mkRel k else EConstr.map_with_binders sigma (fun (k,c) -> (k+1,lift 1 c)) substrec kc t in @@ -775,7 +774,7 @@ let replace_term_gen sigma eq_fun c by_c in_t = EConstr.map_with_binders sigma (fun (k,c) -> (k+1,EConstr.Vars.lift 1 c)) substrec kc t) in - EConstr.Unsafe.to_constr (substrec (0,c) in_t) + substrec (0,c) in_t let replace_term sigma c byc t = replace_term_gen sigma EConstr.eq_constr c byc t @@ -933,12 +932,11 @@ let filtering sigma env cv_pb c1 c2 = aux env cv_pb c1 c2; !evm let decompose_prod_letin sigma c = - let inj c = EConstr.Unsafe.to_constr c in let rec prodec_rec i l sigma c = match EConstr.kind sigma c with - | Prod (n,t,c) -> prodec_rec (succ i) (RelDecl.LocalAssum (n, inj t)::l) sigma c - | LetIn (n,d,t,c) -> prodec_rec (succ i) (RelDecl.LocalDef (n, inj d, inj t)::l) sigma c + | Prod (n,t,c) -> prodec_rec (succ i) (local_assum (n, t)::l) sigma c + | LetIn (n,d,t,c) -> prodec_rec (succ i) (local_def (n, d, t)::l) sigma c | Cast (c,_,_) -> prodec_rec i l sigma c - | _ -> i,l, inj c in + | _ -> i,l, c in prodec_rec 0 [] sigma c (* (nb_lam [na1:T1]...[nan:Tan]c) where c is not an abstraction @@ -969,7 +967,7 @@ let nb_prod_modulo_zeta sigma x = | _ -> n in count 0 x -let align_prod_letin sigma c a : Context.Rel.t * constr = +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"; @@ -980,21 +978,23 @@ let align_prod_letin sigma 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") 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) diff --git a/engine/termops.mli b/engine/termops.mli index 1699d600e..d7d9c743d 100644 --- a/engine/termops.mli +++ b/engine/termops.mli @@ -12,6 +12,7 @@ open Pp open Names open Term +open EConstr open Environ (** printers *) @@ -20,59 +21,56 @@ val pr_sort_family : sorts_family -> std_ppcmds val pr_fix : ('a -> std_ppcmds) -> ('a, 'a) pfixpoint -> 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 set_print_constr : (env -> Constr.constr -> std_ppcmds) -> unit +val print_constr : Constr.constr -> std_ppcmds +val print_constr_env : env -> Constr.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 (** about contexts *) -val push_rel_assum : Name.t * EConstr.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_rel_assum : Name.t * types -> 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 -> Context.Rel.t -> int * Constr.constr option * Constr.types (** 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_list : int -> int -> EConstr.constr list +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 -> EConstr.types -> EConstr.types -val it_mkProd : EConstr.types -> (Name.t * EConstr.types) list -> EConstr.types -val it_mkLambda : EConstr.constr -> (Name.t * EConstr.types) list -> EConstr.constr +val mkProd_wo_LetIn : Context.Rel.Declaration.t -> 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 : EConstr.types -> Context.Rel.t -> EConstr.types -val it_mkLambda_or_LetIn : constr -> Context.Rel.t -> constr -val it_mkNamedProd_or_LetIn : EConstr.types -> Context.Named.t -> EConstr.types -val it_mkNamedProd_wo_LetIn : types -> Context.Named.t -> types -val it_mkNamedLambda_or_LetIn : EConstr.constr -> Context.Named.t -> EConstr.constr +val it_mkProd_wo_LetIn : types -> Context.Rel.t -> types +val it_mkLambda_or_LetIn : Constr.constr -> Context.Rel.t -> Constr.constr +val it_mkNamedProd_or_LetIn : types -> Context.Named.t -> types +val it_mkNamedProd_wo_LetIn : Constr.types -> Context.Named.t -> Constr.types +val it_mkNamedLambda_or_LetIn : constr -> Context.Named.t -> 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 : Evd.evar_map -> (Context.Rel.Declaration.t -> 'a -> 'a) -> - ('a -> EConstr.constr -> EConstr.constr) -> - 'a -> EConstr.constr -> EConstr.constr + ('a -> constr -> constr) -> + 'a -> constr -> constr val map_constr_with_full_binders : Evd.evar_map -> (Context.Rel.Declaration.t -> 'a -> 'a) -> - ('a -> EConstr.t -> EConstr.t) -> 'a -> EConstr.t -> EConstr.t + ('a -> constr -> constr) -> 'a -> constr -> constr (** [fold_constr_with_binders g f n acc c] folds [f n] on the immediate subterms of [c] starting from [acc] and proceeding from left to @@ -81,62 +79,61 @@ 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 : - Evd.evar_map -> ('a -> 'a) -> - ('a -> 'b -> EConstr.t -> 'b) -> 'a -> 'b -> EConstr.t -> 'b +val fold_constr_with_binders : Evd.evar_map -> + ('a -> 'a) -> ('a -> 'b -> constr -> 'b) -> 'a -> 'b -> constr -> 'b -val fold_constr_with_full_binders : - Evd.evar_map -> (Context.Rel.Declaration.t -> 'a -> 'a) -> - ('a -> 'b -> EConstr.t -> 'b) -> 'a -> 'b -> EConstr.t -> 'b +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 : Evd.evar_map -> EConstr.t -> EConstr.t -val drop_extra_implicit_args : Evd.evar_map -> EConstr.t -> EConstr.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 : Evd.evar_map -> EConstr.t -> bool -val occur_existential : Evd.evar_map -> EConstr.t -> bool -val occur_meta_or_existential : Evd.evar_map -> EConstr.t -> bool -val occur_evar : Evd.evar_map -> existential_key -> EConstr.t -> bool -val occur_var : env -> Evd.evar_map -> Id.t -> EConstr.t -> 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 -> Evd.evar_map -> Id.t -> Context.Named.Declaration.t -> bool (** As {!occur_var} but assume the identifier not to be a section variable *) -val local_occur_var : Evd.evar_map -> Id.t -> EConstr.t -> bool +val local_occur_var : Evd.evar_map -> Id.t -> constr -> bool -val free_rels : Evd.evar_map -> EConstr.t -> 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 : Evd.evar_map -> EConstr.t -> EConstr.t -> bool -val dependent_no_evar : Evd.evar_map -> EConstr.t -> EConstr.t -> bool -val dependent_univs : Evd.evar_map -> EConstr.t -> EConstr.t -> bool -val dependent_univs_no_evar : Evd.evar_map -> EConstr.t -> EConstr.t -> bool -val dependent_in_decl : Evd.evar_map -> EConstr.t -> Context.Named.Declaration.t -> bool -val count_occurrences : Evd.evar_map -> EConstr.t -> EConstr.t -> int -val collect_metas : Evd.evar_map -> EConstr.t -> int list -val collect_vars : Evd.evar_map -> EConstr.t -> 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 -> Context.Named.Declaration.t -> 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 : Evd.evar_map -> EConstr.t -> EConstr.t -> bool (** Synonymous of dependent *) +val occur_term : Evd.evar_map -> 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 isMetaOf : Evd.evar_map -> metavariable -> EConstr.constr -> bool +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 : EConstr.t -> constr +val pop : constr -> constr (** {6 ... } *) (** Substitution of an arbitrary large term. Uses equality modulo @@ -145,37 +142,37 @@ val pop : EConstr.t -> constr (** [subst_term_gen eq d c] replaces [d] by [Rel 1] in [c] using [eq] as equality *) val subst_term_gen : Evd.evar_map -> - (Evd.evar_map -> EConstr.t -> EConstr.t -> bool) -> EConstr.t -> EConstr.t -> EConstr.constr + (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 : - Evd.evar_map -> (Evd.evar_map -> EConstr.t -> EConstr.t -> bool) -> - EConstr.t -> EConstr.t -> EConstr.t -> constr + 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 : Evd.evar_map -> EConstr.t -> EConstr.t -> EConstr.constr +val subst_term : Evd.evar_map -> constr -> constr -> constr (** [replace_term d e c] replaces [d] by [e] in [c] *) -val replace_term : Evd.evar_map -> EConstr.t -> EConstr.t -> EConstr.t -> 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 : Evd.evar_map -> (Reduction.conv_pb -> EConstr.constr -> EConstr.constr -> bool) -> - Reduction.conv_pb -> EConstr.constr -> EConstr.constr -> bool -val constr_cmp : Evd.evar_map -> Reduction.conv_pb -> EConstr.constr -> EConstr.constr -> bool -val eq_constr : Evd.evar_map -> EConstr.constr -> EConstr.constr -> bool (* FIXME rename: erases universes*) +val compare_constr_univ : Evd.evar_map -> (Reduction.conv_pb -> constr -> constr -> bool) -> + Reduction.conv_pb -> constr -> constr -> bool +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 : constr -> constr +val eta_reduce_head : Evd.evar_map -> constr -> constr (** Flattens application lists *) -val collapse_appl : Evd.evar_map -> EConstr.t -> constr +val collapse_appl : Evd.evar_map -> constr -> constr -val prod_applist : Evd.evar_map -> EConstr.t -> EConstr.t list -> EConstr.t +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 -> EConstr.t -> EConstr.constr +val strip_outer_cast : Evd.evar_map -> constr -> constr exception CannotFilter @@ -185,32 +182,32 @@ 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 : Evd.evar_map -> Context.Rel.t -> Reduction.conv_pb -> constr -> constr -> subst +type subst = (Context.Rel.t * Constr.constr) Evar.Map.t +val filtering : Evd.evar_map -> Context.Rel.t -> Reduction.conv_pb -> Constr.constr -> Constr.constr -> subst -val decompose_prod_letin : Evd.evar_map -> EConstr.t -> int * Context.Rel.t * constr -val align_prod_letin : Evd.evar_map -> EConstr.t -> EConstr.t -> Context.Rel.t * constr +val decompose_prod_letin : Evd.evar_map -> constr -> int * Context.Rel.t * constr +val align_prod_letin : Evd.evar_map -> constr -> constr -> Context.Rel.t * 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 : Evd.evar_map -> EConstr.t -> int +val nb_lam : Evd.evar_map -> constr -> int (** Similar to [nb_lam], but gives the number of products instead *) -val nb_prod : Evd.evar_map -> EConstr.t -> 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 : Evd.evar_map -> EConstr.t -> 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 : Evd.evar_map -> EConstr.t -> EConstr.constr +val last_arg : Evd.evar_map -> constr -> constr (** Force the decomposition of a term as an applicative one *) -val decompose_app_vect : Evd.evar_map -> EConstr.t -> EConstr.constr * EConstr.constr array +val decompose_app_vect : Evd.evar_map -> constr -> constr * constr array -val adjust_app_list_size : EConstr.constr -> EConstr.constr list -> EConstr.constr -> EConstr.constr list -> - (EConstr.constr * EConstr.constr list * EConstr.constr * EConstr.constr list) -val adjust_app_array_size : EConstr.constr -> EConstr.constr array -> EConstr.constr -> EConstr.constr array -> - (EConstr.constr * EConstr.constr array * EConstr.constr * EConstr.constr array) +val adjust_app_list_size : constr -> constr list -> constr -> constr list -> + (constr * constr list * constr * constr list) +val adjust_app_array_size : constr -> constr array -> constr -> constr array -> + (constr * constr array * constr * constr array) (** name contexts *) type names_context = Name.t list @@ -239,15 +236,15 @@ 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 assums_of_rel_context : Context.Rel.t -> (Name.t * Constr.constr) 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 -> Constr.constr -> Constr.constr) -> Context.Rel.t -> Context.Rel.t val fold_named_context_both_sides : ('a -> Context.Named.Declaration.t -> Context.Named.Declaration.t list -> 'a) -> Context.Named.t -> init:'a -> 'a @@ -256,10 +253,10 @@ val compact_named_context : Context.Named.t -> Context.Compacted.t val clear_named_body : Id.t -> env -> env -val global_vars : env -> Evd.evar_map -> EConstr.t -> Id.t list -val global_vars_set : env -> Evd.evar_map -> EConstr.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 -> Context.Named.Declaration.t -> Id.Set.t -val global_app_of_constr : Evd.evar_map -> EConstr.constr -> Globnames.global_reference puniverses * EConstr.constr option +val global_app_of_constr : Evd.evar_map -> constr -> Globnames.global_reference puniverses * constr option (** Gives an ordered list of hypotheses, closed by dependencies, containing a given set *) @@ -268,15 +265,15 @@ val dependency_closure : env -> Evd.evar_map -> Context.Named.t -> Id.Set.t -> I (** Test if an identifier is the basename of a global reference *) val is_section_variable : Id.t -> bool -val global_of_constr : Evd.evar_map -> EConstr.constr -> Globnames.global_reference puniverses +val global_of_constr : Evd.evar_map -> constr -> Globnames.global_reference puniverses -val is_global : Evd.evar_map -> Globnames.global_reference -> EConstr.constr -> bool +val is_global : Evd.evar_map -> Globnames.global_reference -> constr -> bool -val isGlobalRef : Evd.evar_map -> EConstr.t -> bool +val isGlobalRef : Evd.evar_map -> constr -> bool -val is_template_polymorphic : env -> Evd.evar_map -> EConstr.t -> bool +val is_template_polymorphic : env -> Evd.evar_map -> constr -> bool -val is_Prop : Evd.evar_map -> EConstr.t -> bool +val is_Prop : Evd.evar_map -> constr -> bool (** Combinators on judgments *) @@ -285,5 +282,5 @@ val on_judgment_value : ('c -> 'c) -> ('c, 't) punsafe_judgment -> ('c, 't) puns val on_judgment_type : ('t -> 't) -> ('c, 't) punsafe_judgment -> ('c, 't) punsafe_judgment (** {6 Functions to deal with impossible cases } *) -val set_impossible_default_clause : (unit -> (constr * types) Univ.in_universe_context_set) -> unit +val set_impossible_default_clause : (unit -> (Constr.constr * Constr.types) Univ.in_universe_context_set) -> unit val coq_unit_judge : unit -> EConstr.unsafe_judgment Univ.in_universe_context_set diff --git a/ltac/rewrite.ml b/ltac/rewrite.ml index c92ddf990..a7ff8e142 100644 --- a/ltac/rewrite.ml +++ b/ltac/rewrite.ml @@ -94,7 +94,6 @@ let find_global dir s = fun (evd,cstrs) -> let sigma = Sigma.Unsafe.of_evar_map evd in let Sigma (c, sigma, _) = Evarutil.new_global sigma (Lazy.force gr) in - let c = EConstr.of_constr c in let evd = Sigma.to_evar_map sigma in (evd, cstrs), c @@ -206,7 +205,6 @@ end) = struct fun (evd,cstrs) -> let sigma = Sigma.Unsafe.of_evar_map evd in let Sigma (c, sigma, _) = Evarutil.new_global sigma (Lazy.force l) in - let c = EConstr.of_constr c in let evd = Sigma.to_evar_map sigma in (evd, cstrs), c @@ -215,7 +213,6 @@ end) = struct fun (evd,cstrs) -> let sigma = Sigma.Unsafe.of_evar_map evd in let Sigma (c, sigma, _) = Evarutil.new_global sigma (Lazy.force l) in - let c = EConstr.of_constr c in let evd = Sigma.to_evar_map sigma in (evd, cstrs), c diff --git a/plugins/cc/cctac.ml b/plugins/cc/cctac.ml index 0d48b65d0..7a99c45a8 100644 --- a/plugins/cc/cctac.ml +++ b/plugins/cc/cctac.ml @@ -63,7 +63,6 @@ let rec decompose_term env sigma t= Array.fold_left (fun s t->Appli (s,t)) tf targs | Prod (_,a,_b) when noccurn sigma 1 _b -> let b = Termops.pop _b in - let b = EConstr.of_constr b in let sort_b = sf_of env sigma b in let sort_a = sf_of env sigma a in Appli(Appli(Product (sort_a,sort_b) , @@ -118,7 +117,6 @@ let rec pattern_of_constr env sigma c = List.fold_left Int.Set.union Int.Set.empty lrels | Prod (_,a,_b) when noccurn sigma 1 _b -> let b = Termops.pop _b in - let b = EConstr.of_constr b in let pa,sa = pattern_of_constr env sigma a in let pb,sb = pattern_of_constr env sigma b in let sort_b = sf_of env sigma b in diff --git a/plugins/decl_mode/decl_interp.ml b/plugins/decl_mode/decl_interp.ml index ddf013735..ae057b458 100644 --- a/plugins/decl_mode/decl_interp.ml +++ b/plugins/decl_mode/decl_interp.ml @@ -374,7 +374,7 @@ let interp_cases info env sigma params (pat:cases_pattern_expr) hyps = match st'.st_it with Thesis nam -> {st_it=Thesis nam;st_label=st'.st_label} | This _ -> {st_it = This st.st_it;st_label=st.st_label} in - let thyps = fst (match_hyps blend nam2 (Termops.pop (EConstr.of_constr rest1)) hyps) in + let thyps = fst (match_hyps blend nam2 (Vars.lift (-1) rest1) hyps) in tparams,{pat_vars=tpatvars; pat_aliases=taliases; pat_constr=pat_pat; diff --git a/plugins/firstorder/formula.ml b/plugins/firstorder/formula.ml index 96b991e1f..87bac2fe3 100644 --- a/plugins/firstorder/formula.ml +++ b/plugins/firstorder/formula.ml @@ -76,11 +76,13 @@ type kind_of_formula= | Forall of constr*constr | Atom of constr +let pop t = Vars.lift (-1) t + let kind_of_formula gl term = let normalize=special_nf gl in let cciterm=special_whd gl term in match match_with_imp_term (project gl) (EConstr.of_constr cciterm) with - Some (a,b)-> Arrow(EConstr.Unsafe.to_constr a,(pop b)) + Some (a,b)-> Arrow(EConstr.Unsafe.to_constr a,(pop (EConstr.Unsafe.to_constr b))) |_-> match match_with_forall_term (project gl) (EConstr.of_constr cciterm) with Some (_,a,b)-> Forall(EConstr.Unsafe.to_constr a,EConstr.Unsafe.to_constr b) diff --git a/plugins/firstorder/unify.ml b/plugins/firstorder/unify.ml index 5520c7e35..7cbfb8e7d 100644 --- a/plugins/firstorder/unify.ml +++ b/plugins/firstorder/unify.ml @@ -24,6 +24,8 @@ exception UFAIL of constr*constr let strip_outer_cast t = EConstr.Unsafe.to_constr (strip_outer_cast Evd.empty (EConstr.of_constr t)) (** FIXME *) +let pop t = Vars.lift (-1) t + let unif t1 t2= let evd = Evd.empty in (** FIXME *) let bige=Queue.create () @@ -62,7 +64,7 @@ let unif t1 t2= | Cast(_,_,_),_->Queue.add (strip_outer_cast nt1,nt2) bige | _,Cast(_,_,_)->Queue.add (nt1,strip_outer_cast nt2) bige | (Prod(_,a,b),Prod(_,c,d))|(Lambda(_,a,b),Lambda(_,c,d))-> - Queue.add (a,c) bige;Queue.add (pop (EConstr.of_constr b),pop (EConstr.of_constr d)) bige + Queue.add (a,c) bige;Queue.add (pop b,pop d) bige | Case (_,pa,ca,va),Case (_,pb,cb,vb)-> Queue.add (pa,pb) bige; Queue.add (ca,cb) bige; diff --git a/plugins/funind/functional_principles_proofs.ml b/plugins/funind/functional_principles_proofs.ml index 188368082..cc29d68f5 100644 --- a/plugins/funind/functional_principles_proofs.ml +++ b/plugins/funind/functional_principles_proofs.ml @@ -95,6 +95,7 @@ let list_chop ?(msg="") n l = with Failure (msg') -> failwith (msg ^ msg') +let pop t = Vars.lift (-1) t let make_refl_eq constructor type_of_t t = (* let refl_equal_term = Lazy.force refl_equal in *) @@ -289,7 +290,7 @@ let change_eq env sigma hyp_id (context:Context.Rel.t) x t end_of_type = in let sub = compute_substitution Int.Map.empty (snd t1) (snd t2) in let sub = compute_substitution sub (fst t1) (fst t2) in - let end_of_type_with_pop = Termops.pop (EConstr.of_constr end_of_type) in (*the equation will be removed *) + let end_of_type_with_pop = pop end_of_type in (*the equation will be removed *) let new_end_of_type = (* Ugly hack to prevent Map.fold order change between ocaml-3.08.3 and ocaml-3.08.4 Can be safely replaced by the next comment for Ocaml >= 3.08.4 @@ -311,9 +312,9 @@ let change_eq env sigma hyp_id (context:Context.Rel.t) x t end_of_type = try let witness = Int.Map.find i sub in if is_local_def decl then anomaly (Pp.str "can not redefine a rel!"); - (Termops.pop (EConstr.of_constr end_of_type),ctxt_size,mkLetIn (RelDecl.get_name decl, witness, RelDecl.get_type decl, witness_fun)) + (pop end_of_type,ctxt_size,mkLetIn (RelDecl.get_name decl, witness, RelDecl.get_type decl, witness_fun)) with Not_found -> - (mkProd_or_LetIn decl end_of_type, ctxt_size + 1, mkLambda_or_LetIn decl witness_fun) + (Term.mkProd_or_LetIn decl end_of_type, ctxt_size + 1, mkLambda_or_LetIn decl witness_fun) ) 1 (new_end_of_type,0,witness_fun) @@ -416,7 +417,7 @@ let clean_hyp_with_heq ptes_infos eq_hyps hyp_id env sigma = let coq_I = Coqlib.build_coq_I () in let rec scan_type context type_of_hyp : tactic = if isLetIn type_of_hyp then - let real_type_of_hyp = it_mkProd_or_LetIn type_of_hyp context in + let real_type_of_hyp = Term.it_mkProd_or_LetIn type_of_hyp context in let reduced_type_of_hyp = nf_betaiotazeta real_type_of_hyp in (* length of context didn't change ? *) let new_context,new_typ_of_hyp = @@ -429,13 +430,13 @@ let clean_hyp_with_heq ptes_infos eq_hyps hyp_id env sigma = then begin let (x,t_x,t') = destProd type_of_hyp in - let actual_real_type_of_hyp = it_mkProd_or_LetIn t' context in + let actual_real_type_of_hyp = Term.it_mkProd_or_LetIn t' context in if is_property ptes_infos t_x actual_real_type_of_hyp then begin let pte,pte_args = (destApp t_x) in let (* fix_info *) prove_rec_hyp = (Id.Map.find (destVar pte) ptes_infos).proving_tac in - let popped_t' = Termops.pop (EConstr.of_constr t') in - let real_type_of_hyp = it_mkProd_or_LetIn popped_t' context in + let popped_t' = pop t' in + let real_type_of_hyp = Term.it_mkProd_or_LetIn popped_t' context in let prove_new_type_of_hyp = let context_length = List.length context in tclTHENLIST @@ -486,9 +487,9 @@ let clean_hyp_with_heq ptes_infos eq_hyps hyp_id env sigma = (* observe (str "In "++Ppconstr.pr_id hyp_id++ *) (* str " removing useless precond True" *) (* ); *) - let popped_t' = Termops.pop (EConstr.of_constr t') in + let popped_t' = pop t' in let real_type_of_hyp = - it_mkProd_or_LetIn popped_t' context + Term.it_mkProd_or_LetIn popped_t' context in let prove_trivial = let nb_intro = List.length context in @@ -515,9 +516,9 @@ let clean_hyp_with_heq ptes_infos eq_hyps hyp_id env sigma = ] else if is_trivial_eq t_x then (* t_x := t = t => we remove this precond *) - let popped_t' = Termops.pop (EConstr.of_constr t') in + let popped_t' = pop t' in let real_type_of_hyp = - it_mkProd_or_LetIn popped_t' context + Term.it_mkProd_or_LetIn popped_t' context in let hd,args = destApp t_x in let get_args hd args = @@ -616,8 +617,8 @@ let treat_new_case ptes_infos nb_prod continue_tac term dyn_infos = let fun_body = mkLambda(Anonymous, pf_unsafe_type_of g' term, - Termops.replace_term (project g') term (EConstr.mkRel 1) (EConstr.of_constr dyn_infos.info) - ) + EConstr.Unsafe.to_constr (Termops.replace_term (project g') term (EConstr.mkRel 1) (EConstr.of_constr dyn_infos.info) + )) in let new_body = pf_nf_betaiota g' (EConstr.of_constr (mkApp(fun_body,[| new_term_value |]))) in let new_body = EConstr.Unsafe.to_constr new_body in @@ -988,7 +989,7 @@ let generate_equation_lemma evd fnames f fun_num nb_params nb_args rec_args_num (nb_params + nb_args) t,evd in let eqn = mkApp(Lazy.force eq,[|type_of_f;eq_lhs;eq_rhs|]) in - let lemma_type = it_mkProd_or_LetIn eqn type_ctxt in + let lemma_type = Term.it_mkProd_or_LetIn eqn type_ctxt in (* Pp.msgnl (str "lemma type " ++ Printer.pr_lconstr lemma_type ++ fnl () ++ str "f_body " ++ Printer.pr_lconstr f_body); *) let f_id = Label.to_id (con_label (fst (destConst f))) in let prove_replacement = diff --git a/plugins/funind/functional_principles_types.ml b/plugins/funind/functional_principles_types.ml index b4eb77870..8683f68c6 100644 --- a/plugins/funind/functional_principles_types.ml +++ b/plugins/funind/functional_principles_types.ml @@ -23,6 +23,8 @@ let observe s = if do_observe () then Feedback.msg_debug s +let pop t = Vars.lift (-1) t + (* Transform an inductive induction principle into a functional one @@ -111,7 +113,7 @@ let compute_new_princ_type_from_rel rel_to_fun sorts princ_type = in let dummy_var = mkVar (Id.of_string "________") in let mk_replacement c i args = - let res = mkApp(rel_to_fun.(i), Array.map (fun c -> Termops.pop (EConstr.of_constr c)) (array_get_start args)) in + let res = mkApp(rel_to_fun.(i), Array.map pop (array_get_start args)) in observe (str "replacing " ++ pr_lconstr c ++ str " by " ++ pr_lconstr res); res in @@ -169,25 +171,25 @@ let compute_new_princ_type_from_rel rel_to_fun sorts princ_type = let new_env = Environ.push_rel (LocalAssum (x,t)) env in let new_b,binders_to_remove_from_b = compute_new_princ_type remove new_env b in if List.exists (eq_constr (mkRel 1)) binders_to_remove_from_b - then (Termops.pop (EConstr.of_constr new_b)), filter_map (eq_constr (mkRel 1)) (fun c -> Termops.pop (EConstr.of_constr c)) binders_to_remove_from_b + then (pop new_b), filter_map (eq_constr (mkRel 1)) pop binders_to_remove_from_b else ( bind_fun(new_x,new_t,new_b), list_union_eq eq_constr binders_to_remove_from_t - (List.map (fun c -> Termops.pop (EConstr.of_constr c)) binders_to_remove_from_b) + (List.map pop binders_to_remove_from_b) ) with | Toberemoved -> (* observe (str "Decl of "++Ppconstr.pr_name x ++ str " is removed "); *) let new_b,binders_to_remove_from_b = compute_new_princ_type remove env (substnl [dummy_var] 1 b) in - new_b, List.map (fun c -> Termops.pop (EConstr.of_constr c)) binders_to_remove_from_b + new_b, List.map pop binders_to_remove_from_b | Toberemoved_with_rel (n,c) -> (* observe (str "Decl of "++Ppconstr.pr_name x ++ str " is removed "); *) let new_b,binders_to_remove_from_b = compute_new_princ_type remove env (substnl [c] n b) in - new_b, list_add_set_eq eq_constr (mkRel n) (List.map (fun c -> Termops.pop (EConstr.of_constr c)) binders_to_remove_from_b) + new_b, list_add_set_eq eq_constr (mkRel n) (List.map pop binders_to_remove_from_b) end and compute_new_princ_type_for_letin remove env x v t b = begin @@ -198,25 +200,25 @@ let compute_new_princ_type_from_rel rel_to_fun sorts princ_type = let new_env = Environ.push_rel (LocalDef (x,v,t)) env in let new_b,binders_to_remove_from_b = compute_new_princ_type remove new_env b in if List.exists (eq_constr (mkRel 1)) binders_to_remove_from_b - then (Termops.pop (EConstr.of_constr new_b)),filter_map (eq_constr (mkRel 1)) (fun c -> Termops.pop (EConstr.of_constr c)) binders_to_remove_from_b + then (pop new_b),filter_map (eq_constr (mkRel 1)) pop binders_to_remove_from_b else ( mkLetIn(new_x,new_v,new_t,new_b), list_union_eq eq_constr (list_union_eq eq_constr binders_to_remove_from_t binders_to_remove_from_v) - (List.map (fun c -> Termops.pop (EConstr.of_constr c)) binders_to_remove_from_b) + (List.map pop binders_to_remove_from_b) ) with | Toberemoved -> (* observe (str "Decl of "++Ppconstr.pr_name x ++ str " is removed "); *) let new_b,binders_to_remove_from_b = compute_new_princ_type remove env (substnl [dummy_var] 1 b) in - new_b, List.map (fun c -> Termops.pop (EConstr.of_constr c)) binders_to_remove_from_b + new_b, List.map pop binders_to_remove_from_b | Toberemoved_with_rel (n,c) -> (* observe (str "Decl of "++Ppconstr.pr_name x ++ str " is removed "); *) let new_b,binders_to_remove_from_b = compute_new_princ_type remove env (substnl [c] n b) in - new_b, list_add_set_eq eq_constr (mkRel n) (List.map (fun c -> Termops.pop (EConstr.of_constr c)) binders_to_remove_from_b) + new_b, list_add_set_eq eq_constr (mkRel n) (List.map pop binders_to_remove_from_b) end and compute_new_princ_type_with_acc remove env e (c_acc,to_remove_acc) = let new_e,to_remove_from_e = compute_new_princ_type remove env e diff --git a/plugins/funind/invfun.ml b/plugins/funind/invfun.ml index c02b64c1f..ca066c4cc 100644 --- a/plugins/funind/invfun.ml +++ b/plugins/funind/invfun.ml @@ -399,8 +399,8 @@ let prove_fun_correct evd functional_induction funs_constr graphs_constr schemes match ctxt with | [] | [_] | [_;_] -> anomaly (Pp.str "bad context") | hres::res::decl::ctxt -> - let res = Termops.it_mkLambda_or_LetIn - (Termops.it_mkProd_or_LetIn concl [hres;res]) + let res = Term.it_mkLambda_or_LetIn + (Term.it_mkProd_or_LetIn concl [hres;res]) (LocalAssum (RelDecl.get_name decl, RelDecl.get_type decl) :: ctxt) in res @@ -793,7 +793,7 @@ let derive_correctness make_scheme functional_induction (funs: pconstant list) ( in let type_info = (type_of_lemma_ctxt,type_of_lemma_concl) in graphs_constr.(i) <- graph; - let type_of_lemma = Termops.it_mkProd_or_LetIn type_of_lemma_concl type_of_lemma_ctxt in + let type_of_lemma = Term.it_mkProd_or_LetIn type_of_lemma_concl type_of_lemma_ctxt in let _ = Typing.e_type_of (Global.env ()) evd (EConstr.of_constr type_of_lemma) in let type_of_lemma = nf_zeta (EConstr.of_constr type_of_lemma) in let type_of_lemma = EConstr.Unsafe.to_constr type_of_lemma in @@ -861,7 +861,7 @@ let derive_correctness make_scheme functional_induction (funs: pconstant list) ( let type_info = (type_of_lemma_ctxt,type_of_lemma_concl) in graphs_constr.(i) <- graph; let type_of_lemma = - Termops.it_mkProd_or_LetIn type_of_lemma_concl type_of_lemma_ctxt + Term.it_mkProd_or_LetIn type_of_lemma_concl type_of_lemma_ctxt in let type_of_lemma = nf_zeta (EConstr.of_constr type_of_lemma) in let type_of_lemma = EConstr.Unsafe.to_constr type_of_lemma in diff --git a/plugins/funind/merge.ml b/plugins/funind/merge.ml index 3688b8c15..2840193a9 100644 --- a/plugins/funind/merge.ml +++ b/plugins/funind/merge.ml @@ -32,7 +32,8 @@ module RelDecl = Context.Rel.Declaration (** {2 Useful operations on constr and glob_constr} *) -let rec popn i c = if i<=0 then c else EConstr.of_constr (pop (popn (i-1) c)) +let pop c = Vars.lift (-1) c +let rec popn i c = if i<=0 then c else pop (popn (i-1) c) (** Substitutions in constr *) let compare_constr_nosub t1 t2 = @@ -986,13 +987,13 @@ let relprinctype_to_funprinctype relprinctype nfuns = (* first remove indarg and indarg_in_concl *) let relinfo_noindarg = { relinfo with indarg_in_concl = false; indarg = None; - concl = EConstr.of_constr (remove_last_arg (pop relinfo.concl)); } in + concl = EConstr.of_constr (remove_last_arg (pop (EConstr.Unsafe.to_constr relinfo.concl))); } in (* the nfuns last induction arguments are functional ones: remove them *) let relinfo_argsok = { relinfo_noindarg with nargs = relinfo_noindarg.nargs - nfuns; (* args is in reverse order, so remove fst *) args = remove_n_fst_list nfuns relinfo_noindarg.args; - concl = popn nfuns relinfo_noindarg.concl; + concl = EConstr.of_constr (popn nfuns (EConstr.Unsafe.to_constr relinfo_noindarg.concl)); } in let new_branches = List.map (funify_branches relinfo_argsok nfuns) relinfo_argsok.branches in diff --git a/plugins/funind/recdef.ml b/plugins/funind/recdef.ml index c71174fef..23b308efb 100644 --- a/plugins/funind/recdef.ml +++ b/plugins/funind/recdef.ml @@ -408,6 +408,7 @@ let treat_case forbid_new_ids to_intros finalize_tac nb_lam e infos : tactic = args.(1),args.(2) in let new_b' = Termops.replace_term (project g') (EConstr.of_constr teq_lhs) (EConstr.of_constr teq_rhs) (EConstr.of_constr new_b) in + let new_b' = EConstr.Unsafe.to_constr new_b' in let new_infos = { infos with info = new_b'; @@ -1253,7 +1254,7 @@ let clear_goals = | Prod(Name id as na,t',b) -> let b' = clear_goal b in if noccurn 1 b' && (is_rec_res id) - then Termops.pop (EConstr.of_constr b') + then Vars.lift (-1) b' else if b' == b then t else mkProd(na,t',b') | _ -> Term.map_constr clear_goal t diff --git a/plugins/quote/quote.ml b/plugins/quote/quote.ml index 2ad97c75b..87276f5df 100644 --- a/plugins/quote/quote.ml +++ b/plugins/quote/quote.ml @@ -270,7 +270,7 @@ let compute_ivs f cs gl = (* The Cases predicate is a lambda; we assume no dependency *) let p = match EConstr.kind sigma p with - | Lambda (_,_,p) -> EConstr.of_constr (Termops.pop p) + | Lambda (_,_,p) -> Termops.pop p | _ -> p in diff --git a/plugins/setoid_ring/newring.ml b/plugins/setoid_ring/newring.ml index 63eccaa40..131ecad33 100644 --- a/plugins/setoid_ring/newring.ml +++ b/plugins/setoid_ring/newring.ml @@ -246,7 +246,6 @@ let lapp f args = mkApp(Lazy.force f,args) let plapp evd f args = let fc = Evarutil.e_new_global evd (Lazy.force f) in - let fc = EConstr.of_constr fc in mkApp(fc,args) let dest_rel0 sigma t = @@ -591,7 +590,6 @@ let make_hyp env evd c = let make_hyp_list env evd lH = let carrier = Evarutil.e_new_global evd (Lazy.force coq_hypo) in - let carrier = EConstr.of_constr carrier in let l = List.fold_right (fun c l -> plapp evd coq_cons [|carrier; (make_hyp env evd c); l|]) lH @@ -602,7 +600,6 @@ let make_hyp_list env evd lH = let interp_power env evd pow = let carrier = Evarutil.e_new_global evd (Lazy.force coq_hypo) in - let carrier = EConstr.of_constr carrier in match pow with | None -> let t = ArgArg(Loc.ghost, Lazy.force ltac_inv_morph_nothing) in @@ -618,7 +615,6 @@ let interp_power env evd pow = let interp_sign env evd sign = let carrier = Evarutil.e_new_global evd (Lazy.force coq_hypo) in - let carrier = EConstr.of_constr carrier in match sign with | None -> plapp evd coq_None [|carrier|] | Some spec -> @@ -628,7 +624,6 @@ let interp_sign env evd sign = let interp_div env evd div = let carrier = Evarutil.e_new_global evd (Lazy.force coq_hypo) in - let carrier = EConstr.of_constr carrier in match div with | None -> plapp evd coq_None [|carrier|] | Some spec -> diff --git a/pretyping/cases.ml b/pretyping/cases.ml index c0141f011..01e2db08c 100644 --- a/pretyping/cases.ml +++ b/pretyping/cases.ml @@ -1690,7 +1690,6 @@ let abstract_tycon loc env evdref subst tycon extenv t = (named_context extenv) in let filter = Filter.make (rel_filter @ named_filter) in let candidates = u :: List.map mkRel vl in - let candidates = List.map EConstr.Unsafe.to_constr candidates in let ev = e_new_evar extenv evdref ~src ~filter ~candidates ty in lift k ev in diff --git a/pretyping/detyping.ml b/pretyping/detyping.ml index 87561868f..3d5a5f025 100644 --- a/pretyping/detyping.ml +++ b/pretyping/detyping.ml @@ -210,16 +210,18 @@ let computable p k = && noccur_between 1 (k+1) ccl +let pop t = Vars.lift (-1) t + let lookup_name_as_displayed env t s = let rec lookup avoid n c = match kind_of_term c with | Prod (name,_,c') -> (match compute_displayed_name_in RenamingForGoal avoid name c' with | (Name id,avoid') -> if Id.equal id s then Some n else lookup avoid' (n+1) c' - | (Anonymous,avoid') -> lookup avoid' (n+1) (pop (EConstr.of_constr c'))) + | (Anonymous,avoid') -> lookup avoid' (n+1) (pop c')) | LetIn (name,_,_,c') -> (match compute_displayed_name_in RenamingForGoal avoid name c' with | (Name id,avoid') -> if Id.equal id s then Some n else lookup avoid' (n+1) c' - | (Anonymous,avoid') -> lookup avoid' (n+1) (pop (EConstr.of_constr c'))) + | (Anonymous,avoid') -> lookup avoid' (n+1) (pop c')) | Cast (c,_,_) -> lookup avoid n c | _ -> None in lookup (ids_of_named_context (named_context env)) 1 t @@ -439,7 +441,7 @@ let detype_instance sigma l = else Some (List.map (detype_level sigma) (Array.to_list (Univ.Instance.to_array l))) let rec detype flags avoid env sigma t = - match kind_of_term (collapse_appl sigma (EConstr.of_constr t)) with + match kind_of_term (EConstr.Unsafe.to_constr (collapse_appl sigma (EConstr.of_constr t))) with | Rel n -> (try match lookup_name_of_rel n (fst env) with | Name id -> GVar (dl, id) diff --git a/pretyping/evarconv.ml b/pretyping/evarconv.ml index 9675ae2ea..6dce8627d 100644 --- a/pretyping/evarconv.ml +++ b/pretyping/evarconv.ml @@ -147,7 +147,7 @@ let check_conv_record env sigma (t1,sk1) (t2,sk2) = let _, a, b = destProd sigma t2 in if noccurn sigma 1 b then lookup_canonical_conversion (proji, Prod_cs), - (Stack.append_app [|a;EConstr.of_constr (pop b)|] Stack.empty) + (Stack.append_app [|a;pop b|] Stack.empty) else raise Not_found | Sort s -> lookup_canonical_conversion diff --git a/pretyping/evarsolve.ml b/pretyping/evarsolve.ml index 27436fdd8..3003620d7 100644 --- a/pretyping/evarsolve.ml +++ b/pretyping/evarsolve.ml @@ -189,8 +189,8 @@ let restrict_evar_key evd evk filter candidates = | None -> evar_filter evi | Some filter -> filter in let candidates = match candidates with - | NoUpdate -> evi.evar_candidates - | UpdateWith c -> Some (List.map EConstr.Unsafe.to_constr c) in + | NoUpdate -> Option.map (fun l -> List.map EConstr.of_constr l) evi.evar_candidates + | UpdateWith c -> Some c in let sigma = Sigma.Unsafe.of_evar_map evd in let Sigma (evk, sigma, _) = restrict_evar sigma evk filter candidates in (Sigma.to_evar_map sigma, evk) diff --git a/pretyping/inductiveops.ml b/pretyping/inductiveops.ml index 120adb9fe..1dcd6399e 100644 --- a/pretyping/inductiveops.ml +++ b/pretyping/inductiveops.ml @@ -443,7 +443,7 @@ let build_branch_type env dep p cs = (applist (base,[build_dependent_constructor cs])) cs.cs_args else - it_mkProd_or_LetIn base cs.cs_args + Term.it_mkProd_or_LetIn base cs.cs_args (**************************************************) @@ -575,7 +575,7 @@ let arity_of_case_predicate env (ind,params) dep k = let arsign,_ = get_arity env (ind,params) in let mind = build_dependent_inductive env (ind,params) in let concl = if dep then mkArrow mind (mkSort k) else mkSort k in - it_mkProd_or_LetIn concl arsign + Term.it_mkProd_or_LetIn concl arsign (***********************************************) (* Inferring the sort of parameters of a polymorphic inductive type diff --git a/pretyping/program.ml b/pretyping/program.ml index 8ec6083f7..caa5a5c8a 100644 --- a/pretyping/program.ml +++ b/pretyping/program.ml @@ -31,7 +31,7 @@ let init_reference dir s () = coq_reference "Program" dir s let papp evdref r args = let open EConstr in let gr = delayed_force r in - mkApp (EConstr.of_constr (Evarutil.e_new_global evdref gr), args) + mkApp (Evarutil.e_new_global evdref gr, args) let sig_typ = init_reference ["Init"; "Specif"] "sig" let sig_intro = init_reference ["Init"; "Specif"] "exist" diff --git a/pretyping/reductionops.ml b/pretyping/reductionops.ml index c796a91ca..90c5b241b 100644 --- a/pretyping/reductionops.ml +++ b/pretyping/reductionops.ml @@ -203,10 +203,10 @@ module Cst_stack = struct let reconstruct_head = List.fold_left (fun t (i,args) -> mkApp (t,Array.sub args i (Array.length args - i))) in List.fold_right - (fun (cst,params,args) t -> EConstr.of_constr (Termops.replace_term sigma + (fun (cst,params,args) t -> Termops.replace_term sigma (reconstruct_head d args) (applist (cst, List.rev params)) - t)) cst_l c + t) cst_l c let pr l = let open Pp in @@ -969,7 +969,7 @@ let rec whd_state_gen ?csts ~refold ~tactic_mode flags env sigma = | Rel 1, [] -> let lc = Array.sub cl 0 (napp-1) in let u = if Int.equal napp 1 then f else mkApp (f,lc) in - if noccurn sigma 1 u then (EConstr.of_constr (pop u),Stack.empty),Cst_stack.empty else fold () + if noccurn sigma 1 u then (pop u,Stack.empty),Cst_stack.empty else fold () | _ -> fold () else fold () | _ -> fold ()) @@ -1068,7 +1068,7 @@ let local_whd_state_gen flags sigma = | Rel 1, [] -> let lc = Array.sub cl 0 (napp-1) in let u = if Int.equal napp 1 then f else mkApp (f,lc) in - if noccurn sigma 1 u then (EConstr.of_constr (pop u),Stack.empty) else s + if noccurn sigma 1 u then (pop u,Stack.empty) else s | _ -> s else s | _ -> s) diff --git a/printing/prettyp.ml b/printing/prettyp.ml index e66d3aafe..e2c0f55f8 100644 --- a/printing/prettyp.ml +++ b/printing/prettyp.ml @@ -75,7 +75,7 @@ let print_ref reduce ref = if reduce then let ctx,ccl = Reductionops.splay_prod_assum (Global.env()) Evd.empty (EConstr.of_constr typ) in let ccl = EConstr.Unsafe.to_constr ccl - in it_mkProd_or_LetIn ccl ctx + in Term.it_mkProd_or_LetIn ccl ctx else typ in let univs = Global.universes_of_global ref in let env = Global.env () in diff --git a/stm/lemmas.ml b/stm/lemmas.ml index 04f888a84..9942e911a 100644 --- a/stm/lemmas.ml +++ b/stm/lemmas.ml @@ -466,7 +466,7 @@ let start_proof_com ?inference_hook kind thms hook = evdref := solve_remaining_evars flags env !evdref (Evd.empty,!evdref); let ids = List.map RelDecl.get_name ctx in (compute_proof_name (pi1 kind) sopt, - (nf_evar !evdref (it_mkProd_or_LetIn t' ctx), + (nf_evar !evdref (Term.it_mkProd_or_LetIn t' ctx), (ids, imps @ lift_implicits (List.length ids) imps'), guard))) thms in diff --git a/tactics/autorewrite.ml b/tactics/autorewrite.ml index 029384297..f2e98ee01 100644 --- a/tactics/autorewrite.ml +++ b/tactics/autorewrite.ml @@ -286,7 +286,7 @@ let decompose_applied_relation metas env sigma c ctype left2right = | None -> let ctx,t' = Reductionops.splay_prod_assum env sigma (EConstr.of_constr ctype) in (* Search for underlying eq *) let t' = EConstr.Unsafe.to_constr t' in - match find_rel (it_mkProd_or_LetIn t' ctx) with + match find_rel (Term.it_mkProd_or_LetIn t' ctx) with | Some c -> Some c | None -> None diff --git a/tactics/class_tactics.ml b/tactics/class_tactics.ml index 3a5347bbf..b1d5d8135 100644 --- a/tactics/class_tactics.ml +++ b/tactics/class_tactics.ml @@ -1492,7 +1492,7 @@ let _ = Used in the partial application tactic. *) let rec head_of_constr sigma t = - let t = strip_outer_cast sigma (EConstr.of_constr (collapse_appl sigma t)) in + let t = strip_outer_cast sigma (collapse_appl sigma t) in match EConstr.kind sigma t with | Prod (_,_,c2) -> head_of_constr sigma c2 | LetIn (_,_,_,c2) -> head_of_constr sigma c2 diff --git a/tactics/eqschemes.ml b/tactics/eqschemes.ml index a8ea7446f..e68267584 100644 --- a/tactics/eqschemes.ml +++ b/tactics/eqschemes.ml @@ -77,7 +77,7 @@ let build_dependent_inductive ind (mib,mip) = @ Context.Rel.to_extended_list 0 realargs) let my_it_mkLambda_or_LetIn s c = it_mkLambda_or_LetIn c s -let my_it_mkProd_or_LetIn s c = it_mkProd_or_LetIn c s +let my_it_mkProd_or_LetIn s c = Term.it_mkProd_or_LetIn c s let my_it_mkLambda_or_LetIn_name s c = it_mkLambda_or_LetIn_name (Global.env()) c s diff --git a/tactics/hints.ml b/tactics/hints.ml index 231695c35..d4b73706c 100644 --- a/tactics/hints.ml +++ b/tactics/hints.ml @@ -1211,7 +1211,7 @@ let prepare_hint check (poly,local) env init (sigma,c) = | Evar (evk,args as ev) -> (* We skip the test whether args is the identity or not *) let t = existential_type sigma ev in - let t = List.fold_right (fun (e,id) c -> EConstr.of_constr (replace_term sigma e id c)) !subst t in + let t = List.fold_right (fun (e,id) c -> replace_term sigma e id c) !subst t in if not (closed0 sigma c) then error "Hints with holes dependent on a bound variable not supported."; if occur_existential sigma t then @@ -1225,7 +1225,7 @@ let prepare_hint check (poly,local) env init (sigma,c) = let id = next_ident_away_from default_prepare_hint_ident (fun id -> Id.Set.mem id !vars) in vars := Id.Set.add id !vars; subst := (evar,mkVar id)::!subst; - mkNamedLambda id t (iter (EConstr.of_constr (replace_term sigma evar (mkVar id) c))) in + mkNamedLambda id t (iter (replace_term sigma evar (mkVar id) c)) in let c' = iter c in if check then Pretyping.check_evars (Global.env()) Evd.empty sigma c'; let diff = Univ.ContextSet.diff (Evd.universe_context_set sigma) (Evd.universe_context_set init) in diff --git a/tactics/tactics.ml b/tactics/tactics.ml index 5ee29c089..b2f2797a6 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -4544,7 +4544,7 @@ let induction_gen_l isrec with_evars elim names lc = id_of_name_using_hdchar (Global.env()) (type_of c) Anonymous in let id = new_fresh_id [] x gl in - let newl' = List.map (fun r -> EConstr.of_constr (replace_term sigma c (mkVar id) r)) l' in + let newl' = List.map (fun r -> replace_term sigma c (mkVar id) r) l' in let _ = newlc:=id::!newlc in Tacticals.New.tclTHEN (letin_tac None (Name id) c None allHypsAndConcl) diff --git a/tactics/term_dnet.ml b/tactics/term_dnet.ml index 38342b64d..219abb7fd 100644 --- a/tactics/term_dnet.ml +++ b/tactics/term_dnet.ml @@ -350,9 +350,11 @@ struct TDnet.Idset.fold (fun id acc -> let c_id = Opt.reduce (Ident.constr_of id) in + let c_id = EConstr.of_constr c_id in let (ctx,wc) = - try Termops.align_prod_letin Evd.empty (EConstr.of_constr whole_c) (EConstr.of_constr c_id) (** FIXME *) + try Termops.align_prod_letin Evd.empty (EConstr.of_constr whole_c) c_id (** FIXME *) with Invalid_argument _ -> [],c_id in + let wc = EConstr.Unsafe.to_constr wc in let wc,whole_c = if Opt.direction then whole_c,wc else wc,whole_c in try let _ = Termops.filtering Evd.empty ctx Reduction.CUMUL wc whole_c in diff --git a/toplevel/command.ml b/toplevel/command.ml index 1093b7fdc..820c1b885 100644 --- a/toplevel/command.ml +++ b/toplevel/command.ml @@ -118,7 +118,7 @@ let interp_definition pl bl p red_option c ctypopt = let c, imps2 = interp_casted_constr_evars_impls ~impls env_bl evdref c ty in let nf, subst = Evarutil.e_nf_evars_and_universes evdref in let body = nf (it_mkLambda_or_LetIn c ctx) in - let typ = nf (it_mkProd_or_LetIn ty ctx) in + let typ = nf (Term.it_mkProd_or_LetIn ty ctx) in let beq b1 b2 = if b1 then b2 else not b2 in let impl_eq (x,y,z) (x',y',z') = beq x x' && beq y y' && beq z z' in (* Check that all implicit arguments inferable from the term @@ -583,7 +583,7 @@ let interp_mutual_inductive (paramsl,indl) notations poly prv finite = (* Interpret the arities *) let arities = List.map (interp_ind_arity env_params evdref) indl in - let fullarities = List.map (fun (c, _, _) -> it_mkProd_or_LetIn c ctx_params) arities in + let fullarities = List.map (fun (c, _, _) -> Term.it_mkProd_or_LetIn c ctx_params) arities in let env_ar = push_types env0 indnames fullarities in let env_ar_params = push_rel_context ctx_params env_ar in @@ -850,7 +850,7 @@ let interp_fix_body env_rec evdref impls (_,ctx) fix ccl = let body = interp_casted_constr_evars env evdref ~impls body ccl in it_mkLambda_or_LetIn body ctx) fix.fix_body -let build_fix_type (_,ctx) ccl = it_mkProd_or_LetIn ccl ctx +let build_fix_type (_,ctx) ccl = Term.it_mkProd_or_LetIn ccl ctx let declare_fix ?(opaque = false) (_,poly,_ as kind) pl ctx f ((def,_),eff) t imps = let ce = definition_entry ~opaque ~types:t ~poly ~univs:ctx ~eff def in @@ -946,7 +946,7 @@ let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation = let len = List.length binders_rel in let top_env = push_rel_context binders_rel env in let top_arity = interp_type_evars top_env evdref arityc in - let full_arity = it_mkProd_or_LetIn top_arity binders_rel in + let full_arity = Term.it_mkProd_or_LetIn top_arity binders_rel in let argtyp, letbinders, make = telescope binders_rel in let argname = Id.of_string "recarg" in let arg = LocalAssum (Name argname, argtyp) in @@ -998,7 +998,7 @@ let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation = let intern_arity = substl [projection] top_arity_let in (* substitute the projection of wfarg for something, now intern_arity is in wfarg :: arg *) - let intern_fun_arity_prod = it_mkProd_or_LetIn intern_arity [wfarg 1] in + let intern_fun_arity_prod = Term.it_mkProd_or_LetIn intern_arity [wfarg 1] in let intern_fun_binder = LocalAssum (Name (add_suffix recname "'"), intern_fun_arity_prod) in let curry_fun = let wfpred = mkLambda (Name argid', argtyp, wf_rel_fun (mkRel 1) (mkRel (2 * len + 4))) in @@ -1008,7 +1008,7 @@ let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation = let rcurry = mkApp (rel, [| measure; lift len measure |]) in let lam = LocalAssum (Name (Id.of_string "recproof"), rcurry) in let body = it_mkLambda_or_LetIn app (lam :: binders_rel) in - let ty = it_mkProd_or_LetIn (lift 1 top_arity) (lam :: binders_rel) in + let ty = Term.it_mkProd_or_LetIn (lift 1 top_arity) (lam :: binders_rel) in LocalDef (Name recname, body, ty) in let fun_bl = intern_fun_binder :: [arg] in @@ -1045,7 +1045,7 @@ let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation = let name = add_suffix recname "_func" in let hook l gr _ = let body = it_mkLambda_or_LetIn (mkApp (Universes.constr_of_global gr, [|make|])) binders_rel in - let ty = it_mkProd_or_LetIn top_arity binders_rel in + let ty = Term.it_mkProd_or_LetIn top_arity binders_rel in let pl, univs = Evd.universe_context ?names:pl !evdref in (*FIXME poly? *) let ce = definition_entry ~poly ~types:ty ~univs (Evarutil.nf_evar !evdref body) in @@ -1055,10 +1055,10 @@ let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation = if Impargs.is_implicit_args () || not (List.is_empty impls) then Impargs.declare_manual_implicits false gr [impls] in - let typ = it_mkProd_or_LetIn top_arity binders in + let typ = Term.it_mkProd_or_LetIn top_arity binders in hook, name, typ else - let typ = it_mkProd_or_LetIn top_arity binders_rel in + let typ = Term.it_mkProd_or_LetIn top_arity binders_rel in let hook l gr _ = if Impargs.is_implicit_args () || not (List.is_empty impls) then Impargs.declare_manual_implicits false gr [impls] diff --git a/toplevel/obligations.ml b/toplevel/obligations.ml index 2a0ebf402..6da2f82ab 100644 --- a/toplevel/obligations.ml +++ b/toplevel/obligations.ml @@ -142,10 +142,11 @@ let trunc_named_context n ctx = List.firstn (len - n) ctx let rec chop_product n t = + let pop t = Vars.lift (-1) t in if Int.equal n 0 then Some t else match kind_of_term t with - | Prod (_, _, b) -> if noccurn 1 b then chop_product (pred n) (Termops.pop (EConstr.of_constr b)) else None + | Prod (_, _, b) -> if noccurn 1 b then chop_product (pred n) (pop b) else None | _ -> None let evar_dependencies evm oev = diff --git a/toplevel/record.ml b/toplevel/record.ml index 96221b742..73035deb0 100644 --- a/toplevel/record.ml +++ b/toplevel/record.ml @@ -375,7 +375,7 @@ let structure_signature ctx = let evm = Sigma.to_evar_map evm in let new_tl = Util.List.map_i (fun pos decl -> - RelDecl.map_type (fun t -> Termops.replace_term evm (EConstr.mkRel pos) (EConstr.mkEvar(ev,[||])) (EConstr.of_constr t)) decl) 1 tl in + RelDecl.map_type (fun t -> EConstr.Unsafe.to_constr (Termops.replace_term evm (EConstr.mkRel pos) (EConstr.mkEvar(ev,[||])) (EConstr.of_constr t))) decl) 1 tl in deps_to_evar evm new_tl in deps_to_evar Evd.empty (List.rev ctx) -- cgit v1.2.3 From 531590c223af42c07a93142ab0cea470a98964e6 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Thu, 24 Nov 2016 17:15:15 +0100 Subject: Removing compatibility layers in Retyping --- ltac/extratactics.ml4 | 2 - ltac/rewrite.ml | 21 +-------- ltac/tacinterp.ml | 3 +- plugins/cc/cctac.ml | 4 -- plugins/extraction/extraction.ml | 2 +- plugins/funind/functional_principles_proofs.ml | 1 + plugins/funind/glob_term_to_relation.ml | 1 + plugins/funind/indfun.ml | 1 + plugins/funind/invfun.ml | 2 + plugins/micromega/coq_micromega.ml | 2 +- plugins/setoid_ring/newring.ml | 14 +++--- plugins/ssrmatching/ssrmatching.ml4 | 7 +-- pretyping/cases.ml | 12 ++---- pretyping/classops.ml | 1 + pretyping/coercion.ml | 4 +- pretyping/detyping.ml | 2 +- pretyping/evarconv.ml | 6 +-- pretyping/evarsolve.ml | 59 ++++++++++++++------------ pretyping/evarsolve.mli | 4 +- pretyping/pretyping.ml | 7 +-- pretyping/retyping.ml | 10 ++--- pretyping/retyping.mli | 17 ++++---- pretyping/tacred.ml | 1 - pretyping/typeclasses.ml | 4 +- pretyping/typing.ml | 4 +- pretyping/typing.mli | 4 +- pretyping/unification.ml | 24 +++-------- proofs/clenv.ml | 4 +- proofs/logic.ml | 8 +++- proofs/redexpr.ml | 4 +- proofs/refine.ml | 1 - proofs/tacmach.ml | 4 +- proofs/tacmach.mli | 4 +- tactics/class_tactics.ml | 2 - tactics/eqschemes.ml | 1 + tactics/equality.ml | 13 +----- tactics/hints.ml | 1 - tactics/inv.ml | 4 +- tactics/tactics.ml | 32 +++----------- toplevel/command.ml | 5 ++- toplevel/indschemes.ml | 1 + 41 files changed, 122 insertions(+), 181 deletions(-) diff --git a/ltac/extratactics.ml4 b/ltac/extratactics.ml4 index 138400991..bf8481b52 100644 --- a/ltac/extratactics.ml4 +++ b/ltac/extratactics.ml4 @@ -301,7 +301,6 @@ let project_hint pri l2r r = let sigma, c = Evd.fresh_global env sigma gr in let c = EConstr.of_constr c in let t = Retyping.get_type_of env sigma c in - let t = EConstr.of_constr t in let t = Tacred.reduce_to_quantified_ref env sigma (Lazy.force coq_iff_ref) t in let sign,ccl = decompose_prod_assum sigma t in @@ -667,7 +666,6 @@ let hResolve id c occ t = let t_constr = EConstr.of_constr t_constr in let sigma = Evd.merge_universe_context sigma ctx in let t_constr_type = Retyping.get_type_of env sigma t_constr in - let t_constr_type = EConstr.of_constr t_constr_type in let tac = (change_concl (mkLetIn (Anonymous,t_constr,t_constr_type,concl))) in diff --git a/ltac/rewrite.ml b/ltac/rewrite.ml index a7ff8e142..85c19fac4 100644 --- a/ltac/rewrite.ml +++ b/ltac/rewrite.ml @@ -454,7 +454,7 @@ module TypeGlobal = struct end let sort_of_rel env evm rel = - Reductionops.sort_of_arity env evm (EConstr.of_constr (Retyping.get_type_of env evm rel)) + Reductionops.sort_of_arity env evm (Retyping.get_type_of env evm rel) let is_applied_rewrite_relation = PropGlobal.is_applied_rewrite_relation @@ -524,14 +524,12 @@ let rec decompose_app_rel env evd t = let decompose_app_rel env evd t = let (rel, t1, t2) = decompose_app_rel env evd t in let ty = Retyping.get_type_of env evd rel in - let ty = EConstr.of_constr ty in let () = if not (Reductionops.is_arity env evd ty) then error_no_relation () in (rel, t1, t2) let decompose_applied_relation env sigma (c,l) = let open Context.Rel.Declaration in let ctype = Retyping.get_type_of env sigma c in - let ctype = EConstr.of_constr ctype in let find_rel ty = let sigma, cl = Clenv.make_evar_clause env sigma ty in let sigma = Clenv.solve_evar_clause env sigma true cl l in @@ -539,8 +537,6 @@ let decompose_applied_relation env sigma (c,l) = let (equiv, c1, c2) = decompose_app_rel env sigma t in let ty1 = Retyping.get_type_of env sigma c1 in let ty2 = Retyping.get_type_of env sigma c2 in - let ty1 = EConstr.of_constr ty1 in - let ty2 = EConstr.of_constr ty2 in match evd_convertible env sigma ty1 ty2 with | None -> None | Some sigma -> @@ -750,8 +746,6 @@ let unify_eqn (car, rel, prf, c1, c2, holes, sort) l2r flags env (sigma, cstrs) and prf = nf prf in let ty1 = Retyping.get_type_of env evd c1 in let ty2 = Retyping.get_type_of env evd c2 in - let ty1 = EConstr.of_constr ty1 in - let ty2 = EConstr.of_constr ty2 in let () = if not (convertible env evd ty2 ty1) then raise Reduction.NotConvertible in let rew_evars = evd, cstrs in let rew_prf = RewPrf (rel, prf) in @@ -956,7 +950,6 @@ let reset_env env = let fold_match ?(force=false) env sigma c = let (ci, p, c, brs) = destCase sigma c in let cty = Retyping.get_type_of env sigma c in - let cty = EConstr.of_constr cty in let dep, pred, exists, (sk,eff) = let env', ctx, body = let ctx, pred = decompose_lam_assum sigma p in @@ -1020,7 +1013,6 @@ let subterm all flags (s : 'a pure_strategy) : 'a pure_strategy = state, (None :: acc, evars, progress) else let argty = Retyping.get_type_of env (goalevars evars) arg in - let argty = EConstr.of_constr argty in let state, res = s.strategy { state ; env ; unfresh ; term1 = arg ; ty1 = argty ; @@ -1069,7 +1061,6 @@ let subterm all flags (s : 'a pure_strategy) : 'a pure_strategy = in if flags.on_morphisms then let mty = Retyping.get_type_of env (goalevars evars) m in - let mty = EConstr.of_constr mty in let evars, cstr', m, mty, argsl, args = let argsl = Array.to_list args in let lift = if prop then PropGlobal.lift_cstr else TypeGlobal.lift_cstr in @@ -1113,8 +1104,6 @@ let subterm all flags (s : 'a pure_strategy) : 'a pure_strategy = let b = subst1 mkProp b in let tx = Retyping.get_type_of env (goalevars evars) x and tb = Retyping.get_type_of env (goalevars evars) b in - let tx = EConstr.of_constr tx in - let tb = EConstr.of_constr tb in let arr = if prop then PropGlobal.arrow_morphism else TypeGlobal.arrow_morphism in @@ -1193,7 +1182,6 @@ let subterm all flags (s : 'a pure_strategy) : 'a pure_strategy = let open Context.Rel.Declaration in let env' = Environ.push_rel (local_assum (n', t)) env in let bty = Retyping.get_type_of env' (goalevars evars) b in - let bty = EConstr.of_constr bty in let unlift = if prop then PropGlobal.unlift_cstr else TypeGlobal.unlift_cstr in let state, b' = s.strategy { state ; env = env' ; unfresh ; term1 = b ; ty1 = bty ; @@ -1221,7 +1209,6 @@ let subterm all flags (s : 'a pure_strategy) : 'a pure_strategy = | Case (ci, p, c, brs) -> let cty = Retyping.get_type_of env (goalevars evars) c in - let cty = EConstr.of_constr cty in let evars', eqty = app_poly_sort prop env evars coq_eq [| cty |] in let cstr' = Some eqty in let state, c' = s.strategy { state ; env ; unfresh ; @@ -1500,7 +1487,6 @@ let rewrite_with l2r flags c occs : strategy = { strategy = let apply_strategy (s : strategy) env unfresh concl (prop, cstr) evars = let ty = Retyping.get_type_of env (goalevars evars) concl in - let ty = EConstr.of_constr ty in let _, res = s.strategy { state = () ; env ; unfresh ; term1 = concl ; ty1 = ty ; cstr = (prop, Some cstr) ; evars } in @@ -1917,10 +1903,8 @@ let declare_projection n instance_id r = let sigma,c = Evd.fresh_global env sigma r in let c = EConstr.of_constr c in let ty = Retyping.get_type_of env sigma c in - let ty = EConstr.of_constr ty in let term = proper_projection sigma c ty in let sigma, typ = Typing.type_of env sigma term in - let typ = EConstr.of_constr typ in let ctx, typ = decompose_prod_assum sigma typ in let typ = let n = @@ -2115,7 +2099,7 @@ let unification_rewrite l2r c1 c2 sigma prf car rel but env = and car = nf car and rel = nf rel in check_evar_map_of_evars_defs sigma; let prf = nf prf in - let prfty = nf (EConstr.of_constr (Retyping.get_type_of env sigma prf)) in + let prfty = nf (Retyping.get_type_of env sigma prf) in let sort = sort_of_rel env sigma but in let abs = prf, prfty in let prf = mkRel 1 in @@ -2183,7 +2167,6 @@ let setoid_proof ty fn fallback = try let rel, _, _ = decompose_app_rel env sigma concl in let (sigma, t) = Typing.type_of env sigma rel in - let t = EConstr.of_constr t in let car = snd (List.hd (fst (Reductionops.splay_prod env sigma t))) in (try init_relation_classes () with _ -> raise Not_found); fn env sigma car rel diff --git a/ltac/tacinterp.ml b/ltac/tacinterp.ml index 6c59abe66..b4d2b1e50 100644 --- a/ltac/tacinterp.ml +++ b/ltac/tacinterp.ml @@ -802,7 +802,8 @@ let interp_may_eval f ist env sigma = function (str "Unbound context identifier" ++ pr_id s ++ str".")) | ConstrTypeOf c -> let (sigma,c_interp) = f ist env sigma c in - Typing.type_of ~refresh:true env sigma (EConstr.of_constr c_interp) + let (sigma, t) = Typing.type_of ~refresh:true env sigma (EConstr.of_constr c_interp) in + (sigma, EConstr.Unsafe.to_constr t) | ConstrTerm c -> try f ist env sigma c diff --git a/plugins/cc/cctac.ml b/plugins/cc/cctac.ml index 7a99c45a8..a4ed4798a 100644 --- a/plugins/cc/cctac.ml +++ b/plugins/cc/cctac.ml @@ -264,7 +264,6 @@ let refresh_universes ty k = let env = Proofview.Goal.env gl in let evm = Tacmach.New.project gl in let evm, ty = refresh_type env evm ty in - let ty = EConstr.of_constr ty in Sigma.Unsafe.of_pair (k ty, evm) end } @@ -387,7 +386,6 @@ let discriminate_tac (cstr,u as cstru) p = let trivial = EConstr.of_constr trivial in let evm = Tacmach.New.project gl in let evm, intype = refresh_type env evm (EConstr.of_constr (Tacmach.New.pf_unsafe_type_of gl t1)) in - let intype = EConstr.of_constr intype in let evm, outtype = Evd.new_sort_variable Evd.univ_flexible evm in let outtype = mkSort outtype in let pred = mkLambda(Name xid,outtype,mkRel 1) in @@ -496,9 +494,7 @@ let mk_eq f c1 c2 k = Proofview.Goal.enter { enter = begin fun gl -> let open Tacmach.New in let evm, ty = pf_apply type_of gl c1 in - let ty = EConstr.of_constr ty in let evm, ty = Evarsolve.refresh_universes (Some false) (pf_env gl) evm ty in - let ty = EConstr.of_constr ty in let term = mkApp (fc, [| ty; c1; c2 |]) in let evm, _ = type_of (pf_env gl) evm term in Tacticals.New.tclTHEN (Proofview.V82.tactic (Refiner.tclEVARS evm)) diff --git a/plugins/extraction/extraction.ml b/plugins/extraction/extraction.ml index 61547f96d..7b7e746f2 100644 --- a/plugins/extraction/extraction.ml +++ b/plugins/extraction/extraction.ml @@ -42,7 +42,7 @@ let none = Evd.empty let type_of env c = let polyprop = (lang() == Haskell) in - Retyping.get_type_of ~polyprop env none (strip_outer_cast none (EConstr.of_constr c)) + EConstr.Unsafe.to_constr (Retyping.get_type_of ~polyprop env none (strip_outer_cast none (EConstr.of_constr c))) let sort_of env c = let polyprop = (lang() == Haskell) in diff --git a/plugins/funind/functional_principles_proofs.ml b/plugins/funind/functional_principles_proofs.ml index cc29d68f5..c98cdc467 100644 --- a/plugins/funind/functional_principles_proofs.ml +++ b/plugins/funind/functional_principles_proofs.ml @@ -985,6 +985,7 @@ let generate_equation_lemma evd fnames f fun_num nb_params nb_args rec_args_num let (type_ctxt,type_of_f),evd = let evd,t = Typing.type_of ~refresh:true (Global.env ()) evd (EConstr.of_constr f) in + let t = EConstr.Unsafe.to_constr t in decompose_prod_n_assum (nb_params + nb_args) t,evd in diff --git a/plugins/funind/glob_term_to_relation.ml b/plugins/funind/glob_term_to_relation.ml index 38cd21684..0725bb11c 100644 --- a/plugins/funind/glob_term_to_relation.ml +++ b/plugins/funind/glob_term_to_relation.ml @@ -1273,6 +1273,7 @@ let do_build_inductive Array.fold_right2 (fun id c (evd,env) -> let evd,t = Typing.type_of env evd (EConstr.mkConstU c) in + let t = EConstr.Unsafe.to_constr t in evd, Environ.push_named (LocalAssum (id,t)) (* try *) diff --git a/plugins/funind/indfun.ml b/plugins/funind/indfun.ml index 37a76bec1..1b899c152 100644 --- a/plugins/funind/indfun.ml +++ b/plugins/funind/indfun.ml @@ -373,6 +373,7 @@ let generate_principle (evd:Evd.evar_map ref) pconstants on_error let evd',uprinc = Evd.fresh_global env !evd princ in let _ = evd := evd' in let princ_type = Typing.e_type_of ~refresh:true env evd (EConstr.of_constr uprinc) in + let princ_type = EConstr.Unsafe.to_constr princ_type in Functional_principles_types.generate_functional_principle evd interactive_proof diff --git a/plugins/funind/invfun.ml b/plugins/funind/invfun.ml index ca066c4cc..27528c2dc 100644 --- a/plugins/funind/invfun.ml +++ b/plugins/funind/invfun.ml @@ -135,6 +135,7 @@ let generate_type evd g_to_f f graph i = in evd:=evd'; let graph_arity = Typing.e_type_of (Global.env ()) evd (EConstr.of_constr graph) in + let graph_arity = EConstr.Unsafe.to_constr graph_arity in let ctxt,_ = decompose_prod_assum graph_arity in let fun_ctxt,res_type = match ctxt with @@ -203,6 +204,7 @@ let find_induction_principle evd f = | Some rect_lemma -> let evd',rect_lemma = Evd.fresh_global (Global.env ()) !evd (Globnames.ConstRef rect_lemma) in let evd',typ = Typing.type_of ~refresh:true (Global.env ()) evd' (EConstr.of_constr rect_lemma) in + let typ = EConstr.Unsafe.to_constr typ in evd:=evd'; rect_lemma,typ diff --git a/plugins/micromega/coq_micromega.ml b/plugins/micromega/coq_micromega.ml index f96b189c5..ced572466 100644 --- a/plugins/micromega/coq_micromega.ml +++ b/plugins/micromega/coq_micromega.ml @@ -898,7 +898,7 @@ struct let has_typ gl t1 typ = let ty = Retyping.get_type_of (Tacmach.pf_env gl) (Tacmach.project gl) t1 in - Constr.equal ty typ + EConstr.eq_constr (Tacmach.project gl) ty (EConstr.of_constr typ) let is_convertible gl t1 t2 = diff --git a/plugins/setoid_ring/newring.ml b/plugins/setoid_ring/newring.ml index 131ecad33..c0eeff8d7 100644 --- a/plugins/setoid_ring/newring.ml +++ b/plugins/setoid_ring/newring.ml @@ -344,6 +344,8 @@ let _ = add_map "ring" (****************************************************************************) (* Ring database *) +let pr_constr c = pr_constr (EConstr.Unsafe.to_constr c) + module Cmap = Map.Make(Constr) let from_carrier = Summary.ref Cmap.empty ~name:"ring-tac-carrier-table" @@ -357,12 +359,12 @@ let find_ring_structure env sigma l = let ty = Retyping.get_type_of env sigma t in let check c = let ty' = Retyping.get_type_of env sigma c in - if not (Reductionops.is_conv env sigma (EConstr.of_constr ty) (EConstr.of_constr ty')) then + if not (Reductionops.is_conv env sigma ty ty') then user_err ~hdr:"ring" (str"arguments of ring_simplify do not have all the same type") in List.iter check cl'; - (try ring_for_carrier ty + (try ring_for_carrier (EConstr.to_constr sigma ty) with Not_found -> user_err ~hdr:"ring" (str"cannot find a declared ring structure over"++ @@ -495,7 +497,6 @@ let op_smorph r add mul req m1 m2 = (* (setoid,op_morph) *) let ring_equality env evd (r,add,mul,opp,req) = - let pr_constr c = pr_constr (EConstr.to_constr !evd c) in match EConstr.kind !evd req with | App (f, [| _ |]) when eq_constr_nounivs !evd f (Lazy.force coq_eq) -> let setoid = plapp evd coq_eq_setoid [|r|] in @@ -553,7 +554,6 @@ let build_setoid_params env evd r add mul opp req eqth = let dest_ring env sigma th_spec = let th_typ = Retyping.get_type_of env sigma th_spec in - let th_typ = EConstr.of_constr th_typ in match EConstr.kind sigma th_typ with App(f,[|r;zero;one;add;mul;sub;opp;req|]) when eq_constr_nounivs sigma f (Lazy.force coq_almost_ring_theory) -> @@ -585,7 +585,6 @@ let interp_cst_tac env sigma rk kind (zero,one,add,mul,opp) cst_tac = let make_hyp env evd c = let t = Retyping.get_type_of env !evd c in - let t = EConstr.of_constr t in plapp evd coq_mkhypo [|t;c|] let make_hyp_list env evd lH = @@ -820,7 +819,6 @@ let sf_sr = my_reference"SF_SR" let dest_field env evd th_spec = let open Termops in let th_typ = Retyping.get_type_of env !evd th_spec in - let th_typ = EConstr.of_constr th_typ in match EConstr.kind !evd th_typ with | App(f,[|r;zero;one;add;mul;sub;opp;div;inv;req|]) when is_global !evd (Lazy.force afield_theory) f -> @@ -852,12 +850,12 @@ let find_field_structure env sigma l = let ty = Retyping.get_type_of env sigma t in let check c = let ty' = Retyping.get_type_of env sigma c in - if not (Reductionops.is_conv env sigma (EConstr.of_constr ty) (EConstr.of_constr ty')) then + if not (Reductionops.is_conv env sigma ty ty') then user_err ~hdr:"field" (str"arguments of field_simplify do not have all the same type") in List.iter check cl'; - (try field_for_carrier ty + (try field_for_carrier (EConstr.to_constr sigma ty) with Not_found -> user_err ~hdr:"field" (str"cannot find a declared field structure over"++ diff --git a/plugins/ssrmatching/ssrmatching.ml4 b/plugins/ssrmatching/ssrmatching.ml4 index d34c9325e..9798fa11c 100644 --- a/plugins/ssrmatching/ssrmatching.ml4 +++ b/plugins/ssrmatching/ssrmatching.ml4 @@ -1390,9 +1390,10 @@ let ssrpatterntac _ist (arg_ist,arg) gl = let concl0 = pf_concl gl in let (t, uc), concl_x = fill_occ_pattern (Global.env()) sigma0 concl0 pat noindex 1 in - let gl, tty = pf_type_of gl (EConstr.of_constr t) in - let concl = mkLetIn (Name (id_of_string "selected"), t, tty, concl_x) in - let concl = EConstr.of_constr concl in + let t = EConstr.of_constr t in + let concl_x = EConstr.of_constr concl_x in + let gl, tty = pf_type_of gl t in + let concl = EConstr.mkLetIn (Name (id_of_string "selected"), t, tty, concl_x) in Proofview.V82.of_tactic (convert_concl concl DEFAULTcast) gl (* Register "ssrpattern" tactic *) diff --git a/pretyping/cases.ml b/pretyping/cases.ml index 01e2db08c..565a9725c 100644 --- a/pretyping/cases.ml +++ b/pretyping/cases.ml @@ -1502,7 +1502,7 @@ and compile_alias initial pb (na,orig,(expanded,expanded_typ)) rest = if not (Flags.is_program_mode ()) && (isRel sigma orig || isVar sigma orig) then (* Try to compile first using non expanded alias *) try - if initial then f orig (EConstr.of_constr (Retyping.get_type_of pb.env sigma orig)) + if initial then f orig (Retyping.get_type_of pb.env sigma orig) else just_pop () with e when precatchable_exception e -> (* Try then to compile using expanded alias *) @@ -1517,7 +1517,7 @@ and compile_alias initial pb (na,orig,(expanded,expanded_typ)) rest = (* Could be needed in case of a recursive call which requires to be on a variable for size reasons *) pb.evdref := sigma; - if initial then f orig (EConstr.of_constr (Retyping.get_type_of pb.env !(pb.evdref) orig)) + if initial then f orig (Retyping.get_type_of pb.env !(pb.evdref) orig) else just_pop () @@ -1650,13 +1650,12 @@ let abstract_tycon loc env evdref subst tycon extenv t = | Rel n when is_local_def (lookup_rel n env) -> t | Evar ev -> let ty = get_type_of env !evdref t in - let ty = Evarutil.evd_comb1 (refresh_universes (Some false) env) evdref (EConstr.of_constr ty) in + let ty = Evarutil.evd_comb1 (refresh_universes (Some false) env) evdref ty in let inst = List.map_i (fun i _ -> try list_assoc_in_triple i subst0 with Not_found -> mkRel i) 1 (rel_context env) in - let ty = EConstr.of_constr ty in let ev' = e_new_evar env evdref ~src ty in begin match solve_simple_eqn (evar_conv_x full_transparent_state) env !evdref (None,ev,substl inst ev') with | Success evd -> evdref := evd @@ -1672,10 +1671,8 @@ let abstract_tycon loc env evdref subst tycon extenv t = let vl = List.map pi1 good in let ty = let ty = get_type_of env !evdref t in - let ty = EConstr.of_constr ty in Evarutil.evd_comb1 (refresh_universes (Some false) env) evdref ty in - let ty = EConstr.of_constr ty in let ty = lift (-k) (aux x ty) in let depvl = free_rels !evdref ty in let inst = @@ -1708,7 +1705,6 @@ let build_tycon loc env tycon_env s subst tycon extenv evdref t = | Some t -> let t = abstract_tycon loc tycon_env evdref subst tycon extenv t in let evd,tt = Typing.type_of extenv !evdref t in - let tt = EConstr.of_constr tt in evdref := evd; (t,tt) in let b = e_cumul env evdref tt (mkSort s) (* side effect *) in @@ -2109,7 +2105,6 @@ let constr_of_pat env evdref arsign pat avoid = let app = applist (cstr, List.map (lift (List.length sign)) params) in let app = applist (app, args) in let apptype = Retyping.get_type_of env ( !evdref) app in - let apptype = EConstr.of_constr apptype in let IndType (indf, realargs) = find_rectype env (!evdref) apptype in match alias with Anonymous -> @@ -2370,7 +2365,6 @@ let build_dependent_signature env evdref avoid tomatchs arsign = let t = RelDecl.get_type decl in let t = EConstr.of_constr t in let argt = Retyping.get_type_of env !evdref arg in - let argt = EConstr.of_constr argt in let eq, refl_arg = if Reductionops.is_conv env !evdref argt t then (mk_eq evdref (lift (nargeqs + slift) argt) diff --git a/pretyping/classops.ml b/pretyping/classops.ml index e4331aade..13310c44d 100644 --- a/pretyping/classops.ml +++ b/pretyping/classops.ml @@ -442,6 +442,7 @@ let cache_coercion (_, c) = let it, _ = class_info c.coercion_target in let value, ctx = Universes.fresh_global_instance (Global.env()) c.coercion_type in let typ = Retyping.get_type_of (Global.env ()) Evd.empty (EConstr.of_constr value) in + let typ = EConstr.Unsafe.to_constr typ in let xf = { coe_value = value; coe_type = typ; diff --git a/pretyping/coercion.ml b/pretyping/coercion.ml index 7e8559630..f569d9fc4 100644 --- a/pretyping/coercion.ml +++ b/pretyping/coercion.ml @@ -66,7 +66,7 @@ let apply_coercion_args env evd check isproj argl funj = | h::restl -> (* On devrait pouvoir s'arranger pour qu'on n'ait pas a faire hnf_constr *) match EConstr.kind !evdref (whd_all env !evdref typ) with | Prod (_,c1,c2) -> - if check && not (e_cumul env evdref (EConstr.of_constr (Retyping.get_type_of env !evdref h)) c1) then + if check && not (e_cumul env evdref (Retyping.get_type_of env !evdref h) c1) then raise NoCoercion; apply_rec (h::acc) (subst1 h c2) restl | _ -> anomaly (Pp.str "apply_coercion_args") @@ -498,7 +498,7 @@ let rec inh_conv_coerce_to_fail loc env evd rigidonly v t c1 = let v2 = Option.map (fun v -> beta_applist evd' (lift 1 v,[v1])) v in let t2 = match v2 with | None -> subst_term evd' v1 t2 - | Some v2 -> EConstr.of_constr (Retyping.get_type_of env1 evd' v2) in + | Some v2 -> Retyping.get_type_of env1 evd' v2 in let (evd'',v2') = inh_conv_coerce_to_fail loc env1 evd' rigidonly v2 t2 u2 in (evd'', Option.map (fun v2' -> mkLambda (name, u1, v2')) v2') | _ -> raise (NoCoercionNoUnifier (best_failed_evd,e)) diff --git a/pretyping/detyping.ml b/pretyping/detyping.ml index 3d5a5f025..d4e156fa4 100644 --- a/pretyping/detyping.ml +++ b/pretyping/detyping.ml @@ -504,7 +504,7 @@ let rec detype flags avoid env sigma t = let pb = Environ.lookup_projection p (snd env) in let body = pb.Declarations.proj_body in let ty = Retyping.get_type_of (snd env) sigma (EConstr.of_constr c) in - let ((ind,u), args) = Inductiveops.find_mrectype (snd env) sigma (EConstr.of_constr ty) in + let ((ind,u), args) = Inductiveops.find_mrectype (snd env) sigma ty in let args = List.map EConstr.Unsafe.to_constr args in let body' = strip_lam_assum body in let body' = subst_instance_constr u body' in diff --git a/pretyping/evarconv.ml b/pretyping/evarconv.ml index 6dce8627d..afb0bf6d5 100644 --- a/pretyping/evarconv.ml +++ b/pretyping/evarconv.ml @@ -168,7 +168,7 @@ let check_conv_record env sigma (t1,sk1) (t2,sk2) = | Some c -> (* A primitive projection applied to c *) let ty = Retyping.get_type_of ~lax:true env sigma c in let (i,u), ind_args = - try Inductiveops.find_mrectype env sigma (EConstr.of_constr ty) + try Inductiveops.find_mrectype env sigma ty with _ -> raise Not_found in Stack.append_app_list ind_args Stack.empty, c, sk1 | None -> @@ -882,7 +882,7 @@ and conv_record trs env evd (ctx,(h,h2),c,bs,(params,params1),(us,us2),(sk1,sk2) List.fold_left (fun (i,ks,m,test) b -> if match n with Some n -> Int.equal m n | None -> false then - let ty = EConstr.of_constr (Retyping.get_type_of env i t2) in + let ty = Retyping.get_type_of env i t2 in let test i = evar_conv_x trs env i CUMUL ty (substl ks b) in (i,t2::ks, m-1, test) else @@ -1052,7 +1052,7 @@ let second_order_matching ts env_rhs evd (evk,args) argoccs rhs = let id = NamedDecl.get_id decl' in let t = EConstr.of_constr (NamedDecl.get_type decl') in let evs = ref [] in - let ty = EConstr.of_constr (Retyping.get_type_of env_rhs evd c) in + let ty = Retyping.get_type_of env_rhs evd c in let filter' = filter_possible_projections evd c ty ctxt args in (id,t,c,ty,evs,Filter.make filter',occs) :: make_subst (ctxt',l,occsl) | _, _, [] -> [] diff --git a/pretyping/evarsolve.ml b/pretyping/evarsolve.ml index 3003620d7..de2e46a78 100644 --- a/pretyping/evarsolve.ml +++ b/pretyping/evarsolve.ml @@ -23,6 +23,14 @@ open Evarutil open Pretype_errors open Sigma.Notations +let nlocal_assum (na, t) = + let inj = EConstr.Unsafe.to_constr in + Context.Named.Declaration.LocalAssum (na, inj t) + +let nlocal_def (na, b, t) = + let inj = EConstr.Unsafe.to_constr in + Context.Named.Declaration.LocalDef (na, inj b, inj t) + let normalize_evar evd ev = match EConstr.kind evd (mkEvar ev) with | Evar (evk,args) -> (evk,args) @@ -108,11 +116,11 @@ let refresh_universes ?(status=univ_rigid) ?(onlyalg=false) ?(refreshset=false) | Some dir -> refresh status dir t) else (refresh_term_evars false true t; t) in - if !modified then !evdref, EConstr.Unsafe.to_constr t' else !evdref, EConstr.Unsafe.to_constr t + if !modified then !evdref, t' else !evdref, t let get_type_of_refresh ?(polyprop=true) ?(lax=false) env sigma c = let ty = Retyping.get_type_of ~polyprop ~lax env sigma c in - refresh_universes (Some false) env sigma (EConstr.of_constr ty) + refresh_universes (Some false) env sigma ty (************************) @@ -146,7 +154,7 @@ let recheck_applications conv_algo env evdref t = | App (f, args) -> let () = aux env f in let fty = Retyping.get_type_of env !evdref f in - let argsty = Array.map (fun x -> aux env x; EConstr.of_constr (Retyping.get_type_of env !evdref x)) args in + let argsty = Array.map (fun x -> aux env x; Retyping.get_type_of env !evdref x) args in let rec aux i ty = if i < Array.length argsty then match EConstr.kind !evdref (whd_all env !evdref ty) with @@ -158,7 +166,7 @@ let recheck_applications conv_algo env evdref t = Pretype_errors.error_cannot_unify env evd ~reason (argsty.(i), dom)) | _ -> raise (IllTypedInstance (env, ty, argsty.(i))) else () - in aux 0 (EConstr.of_constr fty) + in aux 0 fty | _ -> iter_with_full_binders !evdref (fun d env -> push_rel d env) aux env t in aux env t @@ -173,7 +181,7 @@ type 'a update = | NoUpdate open Context.Named.Declaration -let inst_of_vars sign = Array.map_of_list (get_id %> EConstr.mkVar) sign +let inst_of_vars sign = Array.map_of_list (get_id %> mkVar) sign let restrict_evar_key evd evk filter candidates = match filter, candidates with @@ -413,9 +421,9 @@ let free_vars_and_rels_up_alias_expansion sigma aliases c = let rec expand_and_check_vars sigma aliases = function | [] -> [] - | a::l when EConstr.isRel sigma a || EConstr.isVar sigma a -> + | a::l when isRel sigma a || isVar sigma a -> let a = expansion_of_var sigma aliases a in - if EConstr.isRel sigma a || EConstr.isVar sigma a then a :: expand_and_check_vars sigma aliases l + if isRel sigma a || isVar sigma a then a :: expand_and_check_vars sigma aliases l else raise Exit | _ -> raise Exit @@ -480,7 +488,7 @@ let is_unification_pattern_meta env evd nb m l t = (* so we need to be a rel <= nb *) if List.for_all (fun x -> isRel evd x && destRel evd x <= nb) l then match find_unification_pattern_args env evd l t with - | Some _ as x when not (dependent evd (EConstr.mkMeta m) t) -> x + | Some _ as x when not (dependent evd (mkMeta m) t) -> x | _ -> None else None @@ -591,15 +599,15 @@ let make_projectable_subst aliases sigma evi args = let define_evar_from_virtual_equation define_fun env evd src t_in_env ty_t_in_sign sign filter inst_in_env = let evd = Sigma.Unsafe.of_evar_map evd in - let Sigma (evar_in_env, evd, _) = new_evar_instance sign evd (EConstr.of_constr ty_t_in_sign) ~filter ~src inst_in_env in + let Sigma (evar_in_env, evd, _) = new_evar_instance sign evd ty_t_in_sign ~filter ~src inst_in_env in let evd = Sigma.to_evar_map evd in let t_in_env = EConstr.of_constr (whd_evar evd (EConstr.Unsafe.to_constr t_in_env)) in let (evk, _) = destEvar evd evar_in_env in - let evd = define_fun env evd None (EConstr.destEvar evd evar_in_env) t_in_env in + let evd = define_fun env evd None (destEvar evd evar_in_env) t_in_env in let ctxt = named_context_of_val sign in let inst_in_sign = inst_of_vars (Filter.filter_list filter ctxt) in let evar_in_sign = mkEvar (evk, inst_in_sign) in - (evd,whd_evar evd (EConstr.Unsafe.to_constr evar_in_sign)) + (evd,EConstr.of_constr (whd_evar evd (EConstr.Unsafe.to_constr evar_in_sign))) (* We have x1..xq |- ?e1 : τ and had to solve something like * Σ; Γ |- ?e1[u1..uq] = (...\y1 ... \yk ... c), where c is typically some @@ -624,7 +632,7 @@ let materialize_evar define_fun env evd k (evk1,args1) ty_in_env = if Evd.is_defined evd evk1 then (* Some circularity somewhere (see e.g. #3209) *) raise MorePreciseOccurCheckNeeeded; - let (evk1,args1) = EConstr.destEvar evd (EConstr.mkEvar (evk1,args1)) in + let (evk1,args1) = destEvar evd (mkEvar (evk1,args1)) in let evi1 = Evd.find_undefined evd evk1 in let env1,rel_sign = env_rel_context_chop k env in let sign1 = evar_hyps evi1 in @@ -641,16 +649,16 @@ let materialize_evar define_fun env evd k (evk1,args1) ty_in_env = let t_in_env = EConstr.of_constr t_in_env in let s = Retyping.get_sort_of env evd t_in_env in let evd,ty_t_in_sign = refresh_universes - ~status:univ_flexible (Some false) env evd (EConstr.mkSort s) in + ~status:univ_flexible (Some false) env evd (mkSort s) in define_evar_from_virtual_equation define_fun env evd src t_in_env ty_t_in_sign sign filter inst_in_env in let evd,d' = match d with - | LocalAssum _ -> evd, Context.Named.Declaration.LocalAssum (id,t_in_sign) + | LocalAssum _ -> evd, nlocal_assum (id,t_in_sign) | LocalDef (_,b,_) -> let b = EConstr.of_constr b in let evd,b = define_evar_from_virtual_equation define_fun env evd src b t_in_sign sign filter inst_in_env in - evd, Context.Named.Declaration.LocalDef (id,b,t_in_sign) in + evd, nlocal_def (id,b,t_in_sign) in (push_named_context_val d' sign, Filter.extend 1 filter, (mkRel 1)::(List.map (lift 1) inst_in_env), (mkRel 1)::(List.map (lift 1) inst_in_sign), @@ -661,11 +669,10 @@ let materialize_evar define_fun env evd k (evk1,args1) ty_in_env = let evd,ev2ty_in_sign = let s = Retyping.get_sort_of env evd ty_in_env in let evd,ty_t_in_sign = refresh_universes - ~status:univ_flexible (Some false) env evd (EConstr.mkSort s) in + ~status:univ_flexible (Some false) env evd (mkSort s) in define_evar_from_virtual_equation define_fun env evd src ty_in_env ty_t_in_sign sign2 filter2 inst2_in_env in let evd = Sigma.Unsafe.of_evar_map evd in - let ev2ty_in_sign = EConstr.of_constr ev2ty_in_sign in let Sigma (ev2_in_sign, evd, _) = new_evar_instance sign2 evd ev2ty_in_sign ~filter:filter2 ~src inst2_in_sign in let evd = Sigma.to_evar_map evd in @@ -899,7 +906,7 @@ let extract_unique_projection = function let extract_candidates sols = try UpdateWith - (List.map (function (id,ProjectVar) -> EConstr.mkVar id | _ -> raise Exit) sols) + (List.map (function (id,ProjectVar) -> mkVar id | _ -> raise Exit) sols) with Exit -> NoUpdate @@ -1171,7 +1178,7 @@ let check_evar_instance evd evk1 body conv_algo = (* FIXME: The body might be ill-typed when this is called from w_merge *) (* This happens in practice, cf MathClasses build failure on 2013-3-15 *) let ty = - try EConstr.of_constr (Retyping.get_type_of ~lax:true evenv evd body) + try Retyping.get_type_of ~lax:true evenv evd body with Retyping.RetypeError _ -> error "Ill-typed evar instance" in match conv_algo evenv evd Reduction.CUMUL ty (EConstr.of_constr evi.evar_concl) with @@ -1378,7 +1385,7 @@ let rec invert_definition conv_algo choose env evd pbty (evk,argsv as ev) rhs = | (id,p)::_::_ -> if choose then (mkVar id, p) else raise (NotUniqueInType sols) in - let ty = lazy (EConstr.of_constr (Retyping.get_type_of env !evdref t)) in + let ty = lazy (Retyping.get_type_of env !evdref t) in let evd = do_projection_effects (evar_define conv_algo ~choose) env ty !evdref p in evdref := evd; c @@ -1440,7 +1447,7 @@ let rec invert_definition conv_algo choose env evd pbty (evk,argsv as ev) rhs = if not !progress then raise (NotEnoughInformationEvarEvar t); (* Make the virtual left evar real *) - let ty = EConstr.of_constr (get_type_of env' evd t) in + let ty = get_type_of env' evd t in let (evd,evar'',ev'') = materialize_evar (evar_define conv_algo ~choose) env' evd k ev ty in (* materialize_evar may instantiate ev' by another evar; adjust it *) @@ -1474,7 +1481,7 @@ let rec invert_definition conv_algo choose env evd pbty (evk,argsv as ev) rhs = | _ -> None with | Some l -> - let ty = EConstr.of_constr (get_type_of env' !evdref t) in + let ty = get_type_of env' !evdref t in let candidates = try let t = @@ -1563,15 +1570,15 @@ and evar_define conv_algo ?(choose=false) env evd pbty (evk,argsv as ev) rhs = str "----> " ++ int ev ++ str " := " ++ print_constr body); raise e in*) - let evd' = check_evar_instance evd' evk (EConstr.of_constr body) conv_algo in - Evd.define evk body evd' + let evd' = check_evar_instance evd' evk body conv_algo in + Evd.define evk (EConstr.Unsafe.to_constr body) evd' with | NotEnoughInformationToProgress sols -> postpone_non_unique_projection env evd pbty ev sols rhs | NotEnoughInformationEvarEvar t -> - add_conv_oriented_pb (pbty,env,EConstr.mkEvar ev,t) evd + add_conv_oriented_pb (pbty,env,mkEvar ev,t) evd | MorePreciseOccurCheckNeeeded -> - add_conv_oriented_pb (pbty,env,EConstr.mkEvar ev,rhs) evd + add_conv_oriented_pb (pbty,env,mkEvar ev,rhs) evd | NotInvertibleUsingOurAlgorithm _ | MetaOccurInBodyInternal as e -> raise e | OccurCheckIn (evd,rhs) -> diff --git a/pretyping/evarsolve.mli b/pretyping/evarsolve.mli index b83147514..f2102f8cd 100644 --- a/pretyping/evarsolve.mli +++ b/pretyping/evarsolve.mli @@ -42,7 +42,7 @@ val refresh_universes : (* Also refresh Prop and Set universes, so that the returned type can be any supertype of the original type *) bool option (* direction: true for levels lower than the existing levels *) -> - env -> evar_map -> types -> evar_map * Constr.types + env -> evar_map -> types -> evar_map * types val solve_refl : ?can_drop:bool -> conv_fun_bool -> env -> evar_map -> bool option -> existential_key -> constr array -> constr array -> evar_map @@ -77,4 +77,4 @@ val remove_instance_local_defs : evar_map -> existential_key -> 'a array -> 'a list val get_type_of_refresh : - ?polyprop:bool -> ?lax:bool -> env -> evar_map -> constr -> evar_map * Constr.types + ?polyprop:bool -> ?lax:bool -> env -> evar_map -> constr -> evar_map * types diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml index 7d2c96bb9..a0d8faab4 100644 --- a/pretyping/pretyping.ml +++ b/pretyping/pretyping.ml @@ -426,7 +426,7 @@ let invert_ltac_bound_name lvar env id0 id = str " which is not bound in current context.") let protected_get_type_of env sigma c = - try EConstr.of_constr (Retyping.get_type_of ~lax:true env.ExtraEnv.env sigma c) + try Retyping.get_type_of ~lax:true env.ExtraEnv.env sigma c with Retyping.RetypeError _ -> user_err (str "Cannot reinterpret " ++ quote (print_constr (EConstr.Unsafe.to_constr c)) ++ @@ -774,9 +774,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre refreshed right away. *) let c = mkApp (f, args) in let c = evd_comb1 (Evarsolve.refresh_universes (Some true) env.ExtraEnv.env) evdref c in - let c = EConstr.of_constr c in let t = Retyping.get_type_of env.ExtraEnv.env !evdref c in - let t = EConstr.of_constr t in make_judge c (* use this for keeping evars: resj.uj_val *) t else resj | _ -> resj @@ -840,7 +838,6 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre let t = evd_comb1 (Evarsolve.refresh_universes ~onlyalg:true ~status:Evd.univ_flexible (Some false) env.ExtraEnv.env) evdref j.uj_type in - let t = EConstr.of_constr t in (* The name specified by ltac is used also to create bindings. So the substitution must also be applied on variables before they are looked up in the rel context. *) @@ -1025,7 +1022,6 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre let tval = evd_comb1 (Evarsolve.refresh_universes ~onlyalg:true ~status:Evd.univ_flexible (Some false) env.ExtraEnv.env) evdref tj.utj_val in - let tval = EConstr.of_constr tval in let tval = nf_evar !evdref tval in let cj, tval = match k with | VMcast -> @@ -1097,7 +1093,6 @@ and pretype_type k0 resolve_tc valcon (env : ExtraEnv.t) evdref lvar = function let s = let sigma = !evdref in let t = Retyping.get_type_of env.ExtraEnv.env sigma v in - let t = EConstr.of_constr t in match EConstr.kind sigma (whd_all env.ExtraEnv.env sigma t) with | Sort s -> s | Evar ev when is_Type (existential_type sigma ev) -> diff --git a/pretyping/retyping.ml b/pretyping/retyping.ml index 7db30bf23..a9529d560 100644 --- a/pretyping/retyping.ml +++ b/pretyping/retyping.ml @@ -109,7 +109,7 @@ let retype ?(polyprop=true) sigma = lift n ty | Var id -> type_of_var env id | Const cst -> EConstr.of_constr (rename_type_of_constant env cst) - | Evar (evk, args) -> EConstr.of_constr (Evd.existential_type sigma (evk, Array.map EConstr.Unsafe.to_constr args)) + | Evar ev -> existential_type sigma ev | Ind ind -> EConstr.of_constr (rename_type_of_inductive env ind) | Construct cstr -> EConstr.of_constr (rename_type_of_constructor env cstr) | Case (_,p,c,lf) -> @@ -210,7 +210,7 @@ let get_sort_of ?(polyprop=true) env sigma t = let get_sort_family_of ?(polyprop=true) env sigma c = let _,_,f,_ = retype ~polyprop sigma in anomaly_on_error (f env) c let type_of_global_reference_knowing_parameters env sigma c args = - let _,_,_,f = retype sigma in EConstr.Unsafe.to_constr (anomaly_on_error (f env c) args) + let _,_,_,f = retype sigma in anomaly_on_error (f env c) args let type_of_global_reference_knowing_conclusion env sigma c conclty = match EConstr.kind sigma c with @@ -238,10 +238,10 @@ let type_of_global_reference_knowing_conclusion env sigma c conclty = let get_type_of ?(polyprop=true) ?(lax=false) env sigma c = let f,_,_,_ = retype ~polyprop sigma in - if lax then EConstr.Unsafe.to_constr (f env c) else EConstr.Unsafe.to_constr (anomaly_on_error (f env) c) + if lax then f env c else anomaly_on_error (f env) c (* Makes an unsafe judgment from a constr *) -let get_judgment_of env evc c = { uj_val = c; uj_type = EConstr.of_constr (get_type_of env evc c) } +let get_judgment_of env evc c = { uj_val = c; uj_type = get_type_of env evc c } (* Returns sorts of a context *) let sorts_of_context env evc ctxt = @@ -256,7 +256,7 @@ let sorts_of_context env evc ctxt = let expand_projection env sigma pr c args = let ty = get_type_of ~lax:true env sigma c in let (i,u), ind_args = - try Inductiveops.find_mrectype env sigma (EConstr.of_constr ty) + try Inductiveops.find_mrectype env sigma ty with Not_found -> retype_error BadRecursiveType in mkApp (mkConstU (Projection.constant pr,u), diff --git a/pretyping/retyping.mli b/pretyping/retyping.mli index a20b11b76..ce9e1635f 100644 --- a/pretyping/retyping.mli +++ b/pretyping/retyping.mli @@ -9,6 +9,7 @@ open Term open Evd open Environ +open EConstr (** This family of functions assumes its constr argument is known to be well-typable. It does not type-check, just recompute the type @@ -26,25 +27,25 @@ type retype_error exception RetypeError of retype_error val get_type_of : - ?polyprop:bool -> ?lax:bool -> env -> evar_map -> EConstr.constr -> types + ?polyprop:bool -> ?lax:bool -> env -> evar_map -> constr -> types val get_sort_of : - ?polyprop:bool -> env -> evar_map -> EConstr.types -> sorts + ?polyprop:bool -> env -> evar_map -> types -> sorts val get_sort_family_of : - ?polyprop:bool -> env -> evar_map -> EConstr.types -> sorts_family + ?polyprop:bool -> env -> evar_map -> types -> sorts_family (** Makes an unsafe judgment from a constr *) -val get_judgment_of : env -> evar_map -> EConstr.constr -> EConstr.unsafe_judgment +val get_judgment_of : env -> evar_map -> constr -> unsafe_judgment -val type_of_global_reference_knowing_parameters : env -> evar_map -> EConstr.constr -> - EConstr.constr array -> types +val type_of_global_reference_knowing_parameters : env -> evar_map -> constr -> + constr array -> types val type_of_global_reference_knowing_conclusion : - env -> evar_map -> EConstr.constr -> EConstr.types -> evar_map * EConstr.types + env -> evar_map -> constr -> types -> evar_map * types val sorts_of_context : env -> evar_map -> Context.Rel.t -> sorts list -val expand_projection : env -> evar_map -> Names.projection -> EConstr.constr -> EConstr.constr list -> EConstr.constr +val expand_projection : env -> evar_map -> Names.projection -> constr -> constr list -> constr val print_retype_error : retype_error -> Pp.std_ppcmds diff --git a/pretyping/tacred.ml b/pretyping/tacred.ml index 02524f896..3fc01c86c 100644 --- a/pretyping/tacred.ml +++ b/pretyping/tacred.ml @@ -1148,7 +1148,6 @@ let compute = cbv_betadeltaiota let abstract_scheme env sigma (locc,a) (c, sigma) = let ta = Retyping.get_type_of env sigma a in - let ta = EConstr.of_constr ta in let na = named_hd env (EConstr.to_constr sigma ta) Anonymous in if occur_meta sigma ta then error "Cannot find a type for the generalisation."; if occur_meta sigma a then diff --git a/pretyping/typeclasses.ml b/pretyping/typeclasses.ml index 9ee34341b..9da7005e0 100644 --- a/pretyping/typeclasses.ml +++ b/pretyping/typeclasses.ml @@ -270,7 +270,7 @@ let add_class cl = let check_instance env sigma c = try let (evd, c) = resolve_one_typeclass env sigma - (EConstr.of_constr (Retyping.get_type_of env sigma c)) in + (Retyping.get_type_of env sigma c) in not (Evd.has_undefined evd) with e when CErrors.noncritical e -> false @@ -314,7 +314,7 @@ let build_subclasses ~check env sigma glob pri = let declare_proj hints (cref, pri, body) = let path' = cref :: path in let ty = Retyping.get_type_of env sigma (EConstr.of_constr body) in - let rest = aux pri body (EConstr.of_constr ty) path' in + let rest = aux pri body ty path' in hints @ (path', pri, body) :: rest in List.fold_left declare_proj [] projs in diff --git a/pretyping/typing.ml b/pretyping/typing.ml index f67e0bddc..d24160ea5 100644 --- a/pretyping/typing.ml +++ b/pretyping/typing.ml @@ -395,7 +395,7 @@ let type_of ?(refresh=false) env evd c = (* side-effect on evdref *) if refresh then Evarsolve.refresh_universes ~onlyalg:true (Some false) env !evdref j.uj_type - else !evdref, EConstr.Unsafe.to_constr j.uj_type + else !evdref, j.uj_type let e_type_of ?(refresh=false) env evdref c = let env = enrich_env env evdref in @@ -405,7 +405,7 @@ let e_type_of ?(refresh=false) env evdref c = let evd, c = Evarsolve.refresh_universes ~onlyalg:true (Some false) env !evdref j.uj_type in let () = evdref := evd in c - else EConstr.Unsafe.to_constr j.uj_type + else j.uj_type let e_solve_evars env evdref c = let env = enrich_env env evdref in diff --git a/pretyping/typing.mli b/pretyping/typing.mli index 94a56b6e1..bf26358a2 100644 --- a/pretyping/typing.mli +++ b/pretyping/typing.mli @@ -19,10 +19,10 @@ val unsafe_type_of : env -> evar_map -> EConstr.constr -> types (** Typecheck a term and return its type + updated evars, optionally refreshing universes *) -val type_of : ?refresh:bool -> env -> evar_map -> EConstr.constr -> evar_map * types +val type_of : ?refresh:bool -> env -> evar_map -> EConstr.constr -> evar_map * EConstr.types (** Variant of [type_of] using references instead of state-passing. *) -val e_type_of : ?refresh:bool -> env -> evar_map ref -> EConstr.constr -> types +val e_type_of : ?refresh:bool -> env -> evar_map ref -> EConstr.constr -> EConstr.types (** Typecheck a type and return its sort *) val e_sort_of : env -> evar_map ref -> EConstr.types -> sorts diff --git a/pretyping/unification.ml b/pretyping/unification.ml index 8a8649f11..233b58e91 100644 --- a/pretyping/unification.ml +++ b/pretyping/unification.ml @@ -122,7 +122,6 @@ let abstract_list_all env evd typ c l = error_cannot_find_well_typed_abstraction env evd p l None | Pretype_errors.PretypeError (env',evd,TypingError x) -> error_cannot_find_well_typed_abstraction env evd p l (Some (env',x)) in - let typp = EConstr.of_constr typp in evd,(p,typp) let set_occurrences_of_last_arg args = @@ -704,7 +703,6 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst : subst0) conv_at_top e (try let tyM = Typing.meta_type sigma k in let tyN = get_type_of curenv ~lax:true sigma cN in - let tyN = EConstr.of_constr tyN in check_compatibility curenv CUMUL flags substn tyN tyM with RetypeError _ -> (* Renounce, maybe metas/evars prevents typing *) sigma) @@ -724,7 +722,6 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst : subst0) conv_at_top e if opt.with_types && flags.check_applied_meta_types then (try let tyM = get_type_of curenv ~lax:true sigma cM in - let tyM = EConstr.of_constr tyM in let tyN = Typing.meta_type sigma k in check_compatibility curenv CUMUL flags substn tyM tyN with RetypeError _ -> @@ -911,8 +908,6 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst : subst0) conv_at_top e try (* Force unification of the types to fill in parameters *) let ty1 = get_type_of curenv ~lax:true sigma c1 in let ty2 = get_type_of curenv ~lax:true sigma c2 in - let ty1 = EConstr.of_constr ty1 in - let ty2 = EConstr.of_constr ty2 in unify_0_with_initial_metas substn true curenv cv_pb { flags with modulo_conv_on_closed_terms = Some full_transparent_state; modulo_delta = full_transparent_state; @@ -978,8 +973,6 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst : subst0) conv_at_top e try (* Ensure we call conversion on terms of the same type *) let tyM = get_type_of curenv ~lax:true sigma m1 in let tyN = get_type_of curenv ~lax:true sigma n1 in - let tyM = EConstr.of_constr tyM in - let tyN = EConstr.of_constr tyN in check_compatibility curenv CUMUL flags substn tyM tyN with RetypeError _ -> (* Renounce, maybe metas/evars prevents typing *) sigma @@ -1267,13 +1260,11 @@ let w_coerce_to_type env evd c cty mvty = let w_coerce env evd mv c = let cty = get_type_of env evd c in let mvty = Typing.meta_type evd mv in - w_coerce_to_type env evd c (EConstr.of_constr cty) mvty + w_coerce_to_type env evd c cty mvty let unify_to_type env sigma flags c status u = let sigma, c = refresh_universes (Some false) env sigma c in - let c = EConstr.of_constr c in let t = get_type_of env sigma (nf_meta sigma c) in - let t = EConstr.of_constr t in let t = nf_betaiota sigma (nf_meta sigma t) in unify_0 env sigma CUMUL flags t u @@ -1406,7 +1397,7 @@ let w_merge env with_types flags (evd,metas,evars : subst0) = let evd' = Sigma.to_evar_map evd' in let (evd'',mc,ec) = unify_0 sp_env evd' CUMUL flags - (EConstr.of_constr (get_type_of sp_env evd' c)) (EConstr.of_constr ev.evar_concl) in + (get_type_of sp_env evd' c) (EConstr.of_constr ev.evar_concl) in let evd''' = w_merge_rec evd'' mc ec [] in if evd' == evd''' then Evd.define sp (EConstr.Unsafe.to_constr c) evd''' @@ -1458,13 +1449,13 @@ let check_types env flags (sigma,_,_ as subst) m n = if isEvar_or_Meta sigma (head_app sigma m) then unify_0_with_initial_metas subst true env CUMUL flags - (EConstr.of_constr (get_type_of env sigma n)) - (EConstr.of_constr (get_type_of env sigma m)) + (get_type_of env sigma n) + (get_type_of env sigma m) else if isEvar_or_Meta sigma (head_app sigma n) then unify_0_with_initial_metas subst true env CUMUL flags - (EConstr.of_constr (get_type_of env sigma m)) - (EConstr.of_constr (get_type_of env sigma n)) + (get_type_of env sigma m) + (get_type_of env sigma n) else subst let try_resolve_typeclasses env evd flag m n = @@ -1595,7 +1586,6 @@ let make_pattern_test from_prefix_of_ind is_correct_type env sigma (pending,c) = else t, [] in let sigma = w_typed_unify env sigma Reduction.CONV flags c t' in let ty = Retyping.get_type_of env sigma t in - let ty = EConstr.of_constr ty in if not (is_correct_type ty) then raise (NotUnifiable None); Some(sigma, t, l2) with @@ -1628,8 +1618,8 @@ let make_eq_test env evd c = let make_abstraction_core name (test,out) env sigma c ty occs check_occs concl = let id = - let ty = Option.map EConstr.Unsafe.to_constr ty in let t = match ty with Some t -> t | None -> get_type_of env sigma c in + let t = EConstr.Unsafe.to_constr t in let x = id_of_name_using_hdchar (Global.env()) t name in let ids = ids_of_named_context (named_context env) in if name == Anonymous then next_ident_away_in_goal x ids else diff --git a/proofs/clenv.ml b/proofs/clenv.ml index a428ab0ed..393e958d3 100644 --- a/proofs/clenv.ml +++ b/proofs/clenv.ml @@ -62,7 +62,7 @@ let refresh_undefined_univs clenv = let clenv_hnf_constr ce t = hnf_constr (cl_env ce) (cl_sigma ce) t -let clenv_get_type_of ce c = EConstr.of_constr (Retyping.get_type_of (cl_env ce) (cl_sigma ce) c) +let clenv_get_type_of ce c = Retyping.get_type_of (cl_env ce) (cl_sigma ce) c exception NotExtensibleClause @@ -673,9 +673,7 @@ let evar_of_binder holes = function let define_with_type sigma env ev c = let open EConstr in let t = Retyping.get_type_of env sigma ev in - let t = EConstr.of_constr t in let ty = Retyping.get_type_of env sigma c in - let ty = EConstr.of_constr ty in let j = Environ.make_judge c ty in let (sigma, j) = Coercion.inh_conv_coerce_to true (Loc.ghost) env sigma j t in let (ev, _) = destEvar sigma ev in diff --git a/proofs/logic.ml b/proofs/logic.ml index 98ad9ebe3..8b890f6f8 100644 --- a/proofs/logic.ml +++ b/proofs/logic.ml @@ -331,7 +331,7 @@ let meta_free_prefix sigma a = let goal_type_of env sigma c = if !check then unsafe_type_of env sigma (EConstr.of_constr c) - else Retyping.get_type_of env sigma (EConstr.of_constr c) + else EConstr.Unsafe.to_constr (Retyping.get_type_of env sigma (EConstr.of_constr c)) let rec mk_refgoals sigma goal goalacc conclty trm = let env = Goal.V82.env sigma goal in @@ -341,6 +341,7 @@ let rec mk_refgoals sigma goal goalacc conclty trm = in if (not !check) && not (occur_meta sigma (EConstr.of_constr trm)) then let t'ty = Retyping.get_type_of env sigma (EConstr.of_constr trm) in + let t'ty = EConstr.Unsafe.to_constr t'ty in let sigma = check_conv_leq_goal env sigma trm t'ty conclty in (goalacc,t'ty,sigma,trm) else @@ -377,6 +378,7 @@ let rec mk_refgoals sigma goal goalacc conclty trm = let args, _ = Option.cata (fun i -> CArray.chop i l) (l, [||]) firstmeta in type_of_global_reference_knowing_parameters env sigma (EConstr.of_constr f) (Array.map EConstr.of_constr args) in + let ty = EConstr.Unsafe.to_constr ty in goalacc, ty, sigma, f else mk_hdgoals sigma goal goalacc f @@ -390,6 +392,7 @@ let rec mk_refgoals sigma goal goalacc conclty trm = let (acc',cty,sigma,c') = mk_hdgoals sigma goal goalacc c in let c = mkProj (p, c') in let ty = get_type_of env sigma (EConstr.of_constr c) in + let ty = EConstr.Unsafe.to_constr ty in (acc',ty,sigma,c) | Case (ci,p,c,lf) -> @@ -439,7 +442,7 @@ and mk_hdgoals sigma goal goalacc trm = if is_template_polymorphic env sigma (EConstr.of_constr f) then let l' = meta_free_prefix sigma l in - (goalacc,type_of_global_reference_knowing_parameters env sigma (EConstr.of_constr f) l',sigma,f) + (goalacc,EConstr.Unsafe.to_constr (type_of_global_reference_knowing_parameters env sigma (EConstr.of_constr f) l'),sigma,f) else mk_hdgoals sigma goal goalacc f in let ((acc'',conclty',sigma), args) = mk_arggoals sigma goal acc' hdty l in @@ -465,6 +468,7 @@ and mk_hdgoals sigma goal goalacc trm = let (acc',cty,sigma,c') = mk_hdgoals sigma goal goalacc c in let c = mkProj (p, c') in let ty = get_type_of env sigma (EConstr.of_constr c) in + let ty = EConstr.Unsafe.to_constr ty in (acc',ty,sigma,c) | _ -> diff --git a/proofs/redexpr.ml b/proofs/redexpr.ml index a830e25d9..8878051c8 100644 --- a/proofs/redexpr.ml +++ b/proofs/redexpr.ml @@ -25,7 +25,7 @@ open Misctypes (* call by value normalisation function using the virtual machine *) let cbv_vm env sigma c = let ctyp = Retyping.get_type_of env sigma c in - Vnorm.cbv_vm env sigma c (EConstr.of_constr ctyp) + Vnorm.cbv_vm env sigma c ctyp let warn_native_compute_disabled = CWarnings.create ~name:"native-compute-disabled" ~category:"native-compiler" @@ -38,7 +38,7 @@ let cbv_native env sigma c = cbv_vm env sigma c) else let ctyp = Retyping.get_type_of env sigma c in - Nativenorm.native_norm env sigma c (EConstr.of_constr ctyp) + Nativenorm.native_norm env sigma c ctyp let whd_cbn flags env sigma t = let (state,_) = diff --git a/proofs/refine.ml b/proofs/refine.ml index 32064aff5..c36bb143e 100644 --- a/proofs/refine.ml +++ b/proofs/refine.ml @@ -135,7 +135,6 @@ let refine ?(unsafe = true) f = let with_type env evd c t = let my_type = Retyping.get_type_of env evd c in - let my_type = EConstr.of_constr my_type in let j = Environ.make_judge c my_type in let (evd,j') = Coercion.inh_conv_coerce_to true (Loc.ghost) env evd j t diff --git a/proofs/tacmach.ml b/proofs/tacmach.ml index d9caadd54..7ffb30fa8 100644 --- a/proofs/tacmach.ml +++ b/proofs/tacmach.ml @@ -108,7 +108,7 @@ let pf_const_value = pf_reduce (fun env _ -> constant_value_in env) let pf_reduce_to_quantified_ind = pf_reduce reduce_to_quantified_ind let pf_reduce_to_atomic_ind = pf_reduce reduce_to_atomic_ind -let pf_hnf_type_of gls = pf_get_type_of gls %> EConstr.of_constr %> pf_whd_all gls +let pf_hnf_type_of gls = pf_get_type_of gls %> pf_whd_all gls let pf_is_matching gl p c = pf_apply Constr_matching.is_matching_conv gl p c let pf_matches gl p c = pf_apply Constr_matching.matches_conv gl p c @@ -234,7 +234,7 @@ module New = struct let pf_hnf_constr gl t = pf_apply hnf_constr gl t let pf_hnf_type_of gl t = - pf_whd_all gl (EConstr.of_constr (pf_get_type_of gl t)) + pf_whd_all gl (pf_get_type_of gl t) let pf_matches gl pat t = pf_apply Constr_matching.matches_conv gl pat t diff --git a/proofs/tacmach.mli b/proofs/tacmach.mli index 06b3e3981..f0471bf66 100644 --- a/proofs/tacmach.mli +++ b/proofs/tacmach.mli @@ -41,7 +41,7 @@ val pf_last_hyp : goal sigma -> Context.Named.Declaration.t val pf_ids_of_hyps : goal sigma -> Id.t list val pf_global : goal sigma -> Id.t -> constr val pf_unsafe_type_of : goal sigma -> EConstr.constr -> types -val pf_type_of : goal sigma -> EConstr.constr -> evar_map * types +val pf_type_of : goal sigma -> EConstr.constr -> evar_map * EConstr.types val pf_hnf_type_of : goal sigma -> EConstr.constr -> EConstr.types val pf_get_hyp : goal sigma -> Id.t -> Context.Named.Declaration.t @@ -109,7 +109,7 @@ module New : sig val pf_concl : ([ `NF ], 'r) Proofview.Goal.t -> EConstr.types val pf_unsafe_type_of : ('a, 'r) Proofview.Goal.t -> EConstr.constr -> Term.types - val pf_type_of : ('a, 'r) Proofview.Goal.t -> EConstr.constr -> evar_map * Term.types + val pf_type_of : ('a, 'r) Proofview.Goal.t -> EConstr.constr -> evar_map * EConstr.types val pf_conv_x : ('a, 'r) Proofview.Goal.t -> EConstr.t -> EConstr.t -> bool val pf_get_new_id : identifier -> ([ `NF ], 'r) Proofview.Goal.t -> identifier diff --git a/tactics/class_tactics.ml b/tactics/class_tactics.ml index b1d5d8135..2f8af6b44 100644 --- a/tactics/class_tactics.ml +++ b/tactics/class_tactics.ml @@ -283,7 +283,6 @@ let clenv_of_prods poly nprods (c, clenv) gl = else let sigma = Tacmach.New.project gl in let ty = Retyping.get_type_of (Proofview.Goal.env gl) sigma c in - let ty = EConstr.of_constr ty in let diff = nb_prod sigma ty - nprods in if Pervasives.(>=) diff 0 then (* Was Some clenv... *) @@ -477,7 +476,6 @@ let pr_depth l = prlist_with_sep (fun () -> str ".") int (List.rev l) let is_Prop env sigma concl = let ty = Retyping.get_type_of env sigma concl in - let ty = EConstr.of_constr ty in match EConstr.kind sigma ty with | Sort (Prop Null) -> true | _ -> false diff --git a/tactics/eqschemes.ml b/tactics/eqschemes.ml index e68267584..855273d3b 100644 --- a/tactics/eqschemes.ml +++ b/tactics/eqschemes.ml @@ -597,6 +597,7 @@ let build_r2l_forward_rew_scheme dep env ind kind = let fix_r2l_forward_rew_scheme (c, ctx') = let t = Retyping.get_type_of (Global.env()) Evd.empty (EConstr.of_constr c) in + let t = EConstr.Unsafe.to_constr t in let ctx,_ = decompose_prod_assum t in match ctx with | hp :: p :: ind :: indargs -> diff --git a/tactics/equality.ml b/tactics/equality.ml index 494f36d7d..e1c39bb34 100644 --- a/tactics/equality.ml +++ b/tactics/equality.ml @@ -188,7 +188,6 @@ let instantiate_lemma_all frzevars gl c ty l l2r concl = let instantiate_lemma gl c ty l l2r concl = let sigma, ct = pf_type_of gl c in - let ct = EConstr.of_constr ct in let t = try snd (reduce_to_quantified_ind (pf_env gl) sigma ct) with UserError _ -> ct in let eqclause = Clenv.make_clenv_binding (pf_env gl) sigma (c,t) l in [eqclause] @@ -452,7 +451,6 @@ let general_rewrite_ebindings_clause cls lft2rgt occs frzevars dep_proof_ok ?tac let sigma = Tacmach.New.project gl in let env = Proofview.Goal.env gl in let ctype = get_type_of env sigma c in - let ctype = EConstr.of_constr ctype in let rels, t = decompose_prod_assum sigma (whd_betaiotazeta sigma ctype) in match match_with_equality_type sigma t with | Some (hdcncl,args) -> (* Fast path: direct leibniz-like rewrite *) @@ -635,8 +633,6 @@ let replace_using_leibniz clause c1 c2 l2r unsafe try_prove_eq_opt = let get_type_of = pf_apply get_type_of gl in let t1 = get_type_of c1 and t2 = get_type_of c2 in - let t1 = EConstr.of_constr t1 in - let t2 = EConstr.of_constr t2 in let evd = if unsafe then Some (Tacmach.New.project gl) else @@ -733,7 +729,6 @@ let _ = let find_positions env sigma t1 t2 = let project env sorts posn t1 t2 = let ty1 = get_type_of env sigma t1 in - let ty1 = EConstr.of_constr ty1 in let s = get_sort_family_of env sigma ty1 in if Sorts.List.mem s sorts then [(List.rev posn,t1,t2)] else [] @@ -856,7 +851,7 @@ let injectable env sigma t1 t2 = let descend_then env sigma head dirn = let IndType (indf,_) = - try find_rectype env sigma (EConstr.of_constr (get_type_of env sigma head)) + try find_rectype env sigma (get_type_of env sigma head) with Not_found -> error "Cannot project on an inductive type derived from a dependency." in let indp,_ = (dest_ind_family indf) in @@ -912,7 +907,6 @@ let build_selector env sigma dirn c ind special default = let (indp,_) = dest_ind_family indf in let ind, _ = check_privacy env indp in let typ = Retyping.get_type_of env sigma default in - let typ = EConstr.of_constr typ in let (mib,mip) = lookup_mind_specif env ind in let deparsign = make_arity_signature env true indf in let p = it_mkLambda_or_LetIn typ deparsign in @@ -932,7 +926,6 @@ let build_coq_I () = EConstr.of_constr (build_coq_I ()) let rec build_discriminator env sigma dirn c = function | [] -> let ind = get_type_of env sigma c in - let ind = EConstr.of_constr ind in let true_0,false_0 = build_coq_True(),build_coq_False() in build_selector env sigma dirn c ind true_0 false_0 @@ -1108,7 +1101,6 @@ let make_tuple env sigma (rterm,rty) lind = assert (not (noccurn sigma lind rty)); let sigdata = find_sigma_data env (get_sort_of env sigma rty) in let sigma, a = type_of ~refresh:true env sigma (mkRel lind) in - let a = EConstr.of_constr a in let na = Context.Rel.Declaration.get_name (lookup_rel lind env) in (* We move [lind] to [1] and lift other rels > [lind] by 1 *) let rty = lift (1-lind) (liftn lind (lind+1) rty) in @@ -1396,7 +1388,6 @@ let inject_at_positions env sigma l2r (eq,_,(t,t1,t2)) eq_clause posns tac = let congr = EConstr.of_constr congr in let pf = applist(congr,[t;resty;injfun;t1;t2]) in let sigma, pf_typ = Typing.type_of env sigma pf in - let pf_typ = EConstr.of_constr pf_typ in let inj_clause = apply_on_clause (pf,pf_typ) eq_clause in let pf = Clenvtac.clenv_value_cast_meta inj_clause in let ty = simplify_args env sigma (clenv_type inj_clause) in @@ -1567,7 +1558,6 @@ let lambda_create env (a,b) = let subst_tuple_term env sigma dep_pair1 dep_pair2 b = let sigma = Sigma.to_evar_map sigma in let typ = get_type_of env sigma dep_pair1 in - let typ = EConstr.of_constr typ in (* We find all possible decompositions *) let decomps1 = decomp_tuple_term env sigma dep_pair1 typ in let decomps2 = decomp_tuple_term env sigma dep_pair2 typ in @@ -1659,7 +1649,6 @@ let cutRewriteInConcl l2r eqn = cutRewriteClause l2r eqn None let substClause l2r c cls = Proofview.Goal.enter { enter = begin fun gl -> let eq = pf_apply get_type_of gl c in - let eq = EConstr.of_constr eq in tclTHENS (cutSubstClause l2r eq cls) [Proofview.tclUNIT (); exact_no_check c] end } diff --git a/tactics/hints.ml b/tactics/hints.ml index d4b73706c..9e9635e8a 100644 --- a/tactics/hints.ml +++ b/tactics/hints.ml @@ -855,7 +855,6 @@ let fresh_global_or_constr env sigma poly cr = let make_resolves env sigma flags pri poly ?name cr = let c, ctx = fresh_global_or_constr env sigma poly cr in let cty = Retyping.get_type_of env sigma c in - let cty = EConstr.of_constr cty in let try_apply f = try Some (f (c, cty, ctx)) with Failure _ -> None in let ents = List.map_filter try_apply diff --git a/tactics/inv.ml b/tactics/inv.ml index e45eb2a16..a398e04dd 100644 --- a/tactics/inv.ml +++ b/tactics/inv.ml @@ -63,10 +63,10 @@ let var_occurs_in_pf gl id = *) -type inversion_status = Dep of EConstr.constr option | NoDep +type inversion_status = Dep of constr option | NoDep let compute_eqn env sigma n i ai = - (mkRel (n-i),EConstr.of_constr (get_type_of env sigma (mkRel (n-i)))) + (mkRel (n-i),get_type_of env sigma (mkRel (n-i))) let make_inv_predicate env evd indf realargs id status concl = let nrealargs = List.length realargs in diff --git a/tactics/tactics.ml b/tactics/tactics.ml index b2f2797a6..574f1c6f3 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -818,13 +818,10 @@ let make_change_arg c pats = let check_types env sigma mayneedglobalcheck deep newc origc = let t1 = Retyping.get_type_of env sigma newc in - let t1 = EConstr.of_constr t1 in if deep then begin let t2 = Retyping.get_type_of env sigma origc in - let t2 = EConstr.of_constr t2 in let sigma, t2 = Evarsolve.refresh_universes ~onlyalg:true (Some false) env sigma t2 in - let t2 = EConstr.of_constr t2 in let sigma, b = infer_conv ~pb:Reduction.CUMUL env sigma t1 t2 in if not b then if @@ -1448,7 +1445,6 @@ let general_elim_clause_gen elimtac indclause elim = let sigma = Tacmach.New.project gl in let (elimc,lbindelimc) = elim.elimbody in let elimt = Retyping.get_type_of env sigma elimc in - let elimt = EConstr.of_constr elimt in let i = match elim.elimindex with None -> index_of_ind_arg sigma elimt | Some i -> i in elimtac elim.elimrename i (elimc, elimt, lbindelimc) indclause @@ -1459,7 +1455,6 @@ let general_elim with_evars clear_flag (c, lbindc) elim = let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in let ct = Retyping.get_type_of env sigma c in - let ct = EConstr.of_constr ct in let t = try snd (reduce_to_quantified_ind env sigma ct) with UserError _ -> ct in let elimtac = elimination_clause_scheme with_evars in let indclause = make_clenv_binding env sigma (c, t) lbindc in @@ -1478,7 +1473,6 @@ let general_case_analysis_in_context with_evars clear_flag (c,lbindc) = let env = Proofview.Goal.env gl in let concl = Proofview.Goal.concl gl in let t = Retyping.get_type_of env (Sigma.to_evar_map sigma) c in - let t = EConstr.of_constr t in let (mind,_) = reduce_to_quantified_ind env (Sigma.to_evar_map sigma) t in let sort = Tacticals.New.elimination_sort_of_goal gl in let Sigma (elim, sigma, p) = @@ -1598,7 +1592,6 @@ let elimination_in_clause_scheme with_evars ?(flags=elim_flags ()) let elimclause' = clenv_fchain ~flags indmv elimclause indclause in let hyp = mkVar id in let hyp_typ = Retyping.get_type_of env sigma hyp in - let hyp_typ = EConstr.of_constr hyp_typ in let hypclause = mk_clenv_from_env env sigma (Some 0) (hyp, hyp_typ) in let elimclause'' = clenv_fchain_in id ~flags hypmv elimclause' hypclause in let new_hyp_typ = clenv_type elimclause'' in @@ -1662,7 +1655,6 @@ let make_projection env sigma params cstr sign elim i n c u = in let app = it_mkLambda_or_LetIn proj sign in let t = Retyping.get_type_of env sigma app in - let t = EConstr.of_constr t in Some (app, t) | None -> None in elim @@ -1673,7 +1665,6 @@ let descend_in_conjunctions avoid tac (err, info) c = let sigma = Tacmach.New.project gl in try let t = Retyping.get_type_of env sigma c in - let t = EConstr.of_constr t in let ((ind,u),t) = reduce_to_quantified_ind env sigma t in let sign,ccl = EConstr.decompose_prod_assum sigma t in match match_with_tuple sigma ccl with @@ -1755,7 +1746,7 @@ let general_apply with_delta with_destruct with_evars clear_flag (loc,(c,lbind : let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in - let thm_ty0 = nf_betaiota sigma (EConstr.of_constr (Retyping.get_type_of env sigma c)) in + let thm_ty0 = nf_betaiota sigma (Retyping.get_type_of env sigma c) in let try_apply thm_ty nprod = try let n = nb_prod_modulo_zeta sigma thm_ty - nprod in @@ -1881,7 +1872,7 @@ let progress_with_clause flags innerclause clause = with Not_found -> error "Unable to unify." let apply_in_once_main flags innerclause env sigma (d,lbind) = - let thm = nf_betaiota sigma (EConstr.of_constr (Retyping.get_type_of env sigma d)) in + let thm = nf_betaiota sigma (Retyping.get_type_of env sigma d) in let rec aux clause = try progress_with_clause flags innerclause clause with e when CErrors.noncritical e -> @@ -1993,7 +1984,6 @@ let exact_check c = let env = Proofview.Goal.env gl in let sigma = Sigma.to_evar_map sigma in let sigma, ct = Typing.type_of env sigma c in - let ct = EConstr.of_constr ct in let tac = Tacticals.New.tclTHEN (convert_leq ct concl) (exact_no_check c) in @@ -2662,9 +2652,7 @@ let letin_tac_gen with_eq (id,depdecls,lastlhyp,ccl,c) ty = | Some t -> Sigma.here t sigma | None -> let t = typ_of env sigma c in - let t = EConstr.of_constr t in let sigma, c = Evarsolve.refresh_universes ~onlyalg:true (Some false) env (Sigma.to_evar_map sigma) t in - let c = EConstr.of_constr c in Sigma.Unsafe.of_pair (c, sigma) in let Sigma ((newcl, eq_tac), sigma, q) = match with_eq with @@ -2717,7 +2705,7 @@ let insert_before decls lasthyp env = let mkletin_goal env sigma store with_eq dep (id,lastlhyp,ccl,c) ty = let open Context.Named.Declaration in - let t = match ty with Some t -> t | _ -> EConstr.of_constr (typ_of env sigma c) in + let t = match ty with Some t -> t | _ -> typ_of env sigma c in let decl = if dep then nlocal_def (id,c,t) else nlocal_assum (id,t) in @@ -2850,7 +2838,6 @@ let generalize_goal gl i ((occs,c,b),na as o) (cl,sigma) = let env = Tacmach.pf_env gl in let ids = Tacmach.pf_ids_of_hyps gl in let sigma, t = Typing.type_of env sigma c in - let t = EConstr.of_constr t in generalize_goal_gen env sigma ids i o t cl let old_generalize_dep ?(with_let=false) c gl = @@ -2923,7 +2910,6 @@ let new_generalize_gen_let lconstr = List.fold_right_i (fun i ((_,c,b),_ as o) (cl, sigma, args) -> let sigma, t = Typing.type_of env sigma c in - let t = EConstr.of_constr t in let args = if Option.is_empty b then c :: args else args in let cl, sigma = generalize_goal_gen env sigma ids i o t cl in (cl, sigma, args)) @@ -2974,7 +2960,7 @@ let specialize (c,lbind) ipat = let sigma = Typeclasses.resolve_typeclasses env sigma in sigma, nf_evar sigma c else - let clause = make_clenv_binding env sigma (c,EConstr.of_constr (Retyping.get_type_of env sigma c)) lbind in + let clause = make_clenv_binding env sigma (c,Retyping.get_type_of env sigma c) lbind in let flags = { (default_unify_flags ()) with resolve_evars = true } in let clause = clenv_unify_meta_types ~flags clause in let (thd,tstack) = whd_nored_stack clause.evd (clenv_value clause) in @@ -2991,7 +2977,6 @@ let specialize (c,lbind) ipat = str "."); clause.evd, term in let typ = Retyping.get_type_of env sigma term in - let typ = EConstr.of_constr typ in let tac = match EConstr.kind sigma (fst(EConstr.decompose_app sigma (snd(EConstr.decompose_lam_assum sigma c)))) with | Var id when Id.List.mem id (Tacmach.New.pf_ids_of_hyps gl) -> @@ -3699,7 +3684,6 @@ let abstract_args gl generalize_vars dep id defined f args = let argty = EConstr.of_constr argty in let sigma', ty = Evarsolve.refresh_universes (Some true) env !sigma ty in let () = sigma := sigma' in - let ty = EConstr.of_constr ty in let lenctx = List.length ctx in let liftargty = lift lenctx argty in let leq = constr_cmp !sigma Reduction.CUMUL liftargty ty in @@ -3751,7 +3735,7 @@ let abstract_args gl generalize_vars dep id defined f args = else [] in let body, c' = - if defined then Some c', EConstr.of_constr (Retyping.get_type_of ctxenv !sigma c') + if defined then Some c', Retyping.get_type_of ctxenv !sigma c' else None, c' in let typ = Tacmach.pf_get_hyp_typ gl id in @@ -4339,7 +4323,6 @@ let clear_unselected_context id inhyps cls = let use_bindings env sigma elim must_be_closed (c,lbind) typ = let sigma = Sigma.to_evar_map sigma in - let typ = EConstr.of_constr typ in let typ = if elim == None then (* w/o an scheme, the term has to be applied at least until @@ -4389,7 +4372,6 @@ let check_enough_applied env sigma elim = let t,_ = decompose_app sigma (whd_all env sigma u) in isInd sigma t | Some elimc -> let elimt = Retyping.get_type_of env sigma (fst elimc) in - let elimt = EConstr.of_constr elimt in let scheme = compute_elim_sig sigma ~elimc elimt in match scheme.indref with | None -> @@ -4435,7 +4417,6 @@ let pose_induction_arg_then isrec with_evars (is_arg_pure_hyp,from_prefix) elim let b = not with_evars && with_eq != None in let Sigma (c, sigma, p) = use_bindings env sigma elim b (c0,lbind) t0 in let t = Retyping.get_type_of env (Sigma.to_evar_map sigma) c in - let t = EConstr.of_constr t in let Sigma (ans, sigma, q) = mkletin_goal env sigma store with_eq false (id,lastlhyp,ccl,c) (Some t) in Sigma (ans, sigma, p +> q) end }; @@ -4487,7 +4468,7 @@ let induction_gen clear_flag isrec with_evars elim && lbind == NoBindings && not with_evars && Option.is_empty eqname && clear_flag == None && has_generic_occurrences_but_goal cls (destVar evd c) env evd ccl in - let enough_applied = check_enough_applied env sigma elim (EConstr.of_constr t) in + let enough_applied = check_enough_applied env sigma elim t in if is_arg_pure_hyp && enough_applied then (* First case: induction on a variable already in an inductive type and with maximal abstraction over the variable. @@ -4504,6 +4485,7 @@ let induction_gen clear_flag isrec with_evars elim declaring the induction argument as a new local variable *) let id = (* Type not the right one if partially applied but anyway for internal use*) + let t = EConstr.Unsafe.to_constr t in let x = id_of_name_using_hdchar (Global.env()) t Anonymous in new_fresh_id [] x gl in let info_arg = (is_arg_pure_hyp, not enough_applied) in diff --git a/toplevel/command.ml b/toplevel/command.ml index 820c1b885..1fc89b8b0 100644 --- a/toplevel/command.ml +++ b/toplevel/command.ml @@ -206,7 +206,7 @@ let do_definition ident k pl bl red_option c ctypopt hook = assert(Univ.ContextSet.is_empty ctx); let typ = match ce.const_entry_type with | Some t -> t - | None -> Retyping.get_type_of env evd (EConstr.of_constr c) + | None -> EConstr.Unsafe.to_constr (Retyping.get_type_of env evd (EConstr.of_constr c)) in Obligations.check_evars env evd; let obls, _, c, cty = @@ -459,7 +459,7 @@ let sign_level env evd sign = | LocalDef _ -> lev, push_rel d env | LocalAssum _ -> let s = destSort (Reduction.whd_all env - (nf_evar evd (Retyping.get_type_of env evd (EConstr.of_constr (RelDecl.get_type d))))) + (nf_evar evd (EConstr.Unsafe.to_constr (Retyping.get_type_of env evd (EConstr.of_constr (RelDecl.get_type d)))))) in let u = univ_of_sort s in (Univ.sup u lev, push_rel d env)) @@ -1106,6 +1106,7 @@ let interp_recursive isfix fixl notations = (fun env' id t -> if Flags.is_program_mode () then let sort = Evarutil.evd_comb1 (Typing.type_of ~refresh:true env) evdref (EConstr.of_constr t) in + let sort = EConstr.Unsafe.to_constr sort in let fixprot = try let app = mkApp (delayed_force fix_proto, [|sort; t|]) in diff --git a/toplevel/indschemes.ml b/toplevel/indschemes.ml index 9ed34d713..21bc895a8 100644 --- a/toplevel/indschemes.ml +++ b/toplevel/indschemes.ml @@ -404,6 +404,7 @@ let do_mutual_induction_scheme lnamedepindsort = let sigma, listdecl = Indrec.build_mutual_induction_scheme env0 sigma lrecspec in let declare decl fi lrecref = let decltype = Retyping.get_type_of env0 sigma (EConstr.of_constr decl) in + let decltype = EConstr.to_constr sigma decltype in let proof_output = Future.from_val ((decl,Univ.ContextSet.empty),Safe_typing.empty_private_constants) in let cst = define fi UserIndividualRequest sigma proof_output (Some decltype) in ConstRef cst :: lrecref -- cgit v1.2.3 From 05afd04095e35d77ca135bd2c1cb8d303ea2d6a8 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Thu, 24 Nov 2016 18:18:17 +0100 Subject: Ltac now uses evar-based constrs. --- dev/top_printers.ml | 14 ++-- engine/evd.ml | 6 +- engine/evd.mli | 2 +- engine/termops.ml | 18 ++--- engine/termops.mli | 6 +- interp/stdarg.mli | 5 +- ltac/coretactics.ml4 | 20 +++--- ltac/evar_tactics.ml | 5 +- ltac/evar_tactics.mli | 2 +- ltac/extraargs.mli | 4 +- ltac/extratactics.ml4 | 61 ++++++++--------- ltac/g_auto.ml4 | 12 ++-- ltac/g_class.ml4 | 8 +-- ltac/g_eqdecide.ml4 | 2 +- ltac/g_rewrite.ml4 | 2 +- ltac/pptactic.ml | 32 ++++----- ltac/pptactic.mli | 4 +- ltac/pptacticsig.mli | 2 +- ltac/taccoerce.ml | 57 ++++++++-------- ltac/taccoerce.mli | 25 +++---- ltac/tacinterp.ml | 130 ++++++++++++++++-------------------- ltac/tacinterp.mli | 1 + ltac/tactic_debug.ml | 15 ++--- ltac/tactic_debug.mli | 7 +- ltac/tauto.ml | 2 +- plugins/cc/ccalgo.ml | 8 +-- plugins/cc/cctac.ml | 3 +- plugins/cc/g_congruence.ml4 | 4 +- plugins/decl_mode/ppdecl_proof.ml | 2 + plugins/funind/g_indfun.ml4 | 8 ++- plugins/nsatz/g_nsatz.ml4 | 2 +- plugins/quote/g_quote.ml4 | 5 +- plugins/setoid_ring/g_newring.ml4 | 14 +--- plugins/setoid_ring/newring.ml | 5 +- plugins/ssrmatching/ssrmatching.ml4 | 2 +- pretyping/evardefine.ml | 4 +- pretyping/evardefine.mli | 2 +- pretyping/pretyping.ml | 3 +- pretyping/pretyping.mli | 2 +- pretyping/recordops.ml | 6 +- pretyping/reductionops.ml | 6 +- pretyping/unification.ml | 5 +- printing/printer.ml | 16 +++-- printing/printer.mli | 4 ++ proofs/clenv.ml | 5 +- proofs/redexpr.ml | 13 ++-- proofs/redexpr.mli | 1 + proofs/tacmach.ml | 2 +- tactics/auto.ml | 2 +- tactics/tactics.ml | 4 +- 50 files changed, 282 insertions(+), 288 deletions(-) diff --git a/dev/top_printers.ml b/dev/top_printers.ml index 8290ca945..04aacdc09 100644 --- a/dev/top_printers.ml +++ b/dev/top_printers.ml @@ -70,10 +70,10 @@ let ppwf_paths x = pp (Rtree.pp_tree pprecarg x) (* term printers *) let rawdebug = ref false let ppevar evk = pp (str (Evd.string_of_existential evk)) -let ppconstr x = pp (Termops.print_constr x) -let ppeconstr x = pp (Termops.print_constr (EConstr.Unsafe.to_constr x)) +let ppconstr x = pp (Termops.print_constr (EConstr.of_constr x)) +let ppeconstr x = pp (Termops.print_constr x) let ppconstr_expr x = pp (Ppconstr.pr_constr_expr x) -let ppconstrdb x = pp(Flags.with_option rawdebug Termops.print_constr x) +let ppconstrdb x = pp(Flags.with_option rawdebug Termops.print_constr (EConstr.of_constr x)) let ppterm = ppconstr let ppsconstr x = ppconstr (Mod_subst.force_constr x) let ppconstr_univ x = Constrextern.with_universes ppconstr x @@ -98,9 +98,9 @@ let ppidmap pr l = pp (pridmap pr l) let ppevarsubst = ppidmap (fun id0 -> prset (fun (c,copt,id) -> hov 0 - (Termops.print_constr c ++ + (Termops.print_constr (EConstr.of_constr c) ++ (match copt with None -> mt () | Some c -> spc () ++ str "") ++ + Termops.print_constr (EConstr.of_constr c) ++ str">") ++ (if id = id0 then mt () else spc () ++ str "")))) @@ -109,7 +109,7 @@ let ppididmap = ppidmap (fun _ -> pr_id) let prconstrunderbindersidmap = pridmap (fun _ (l,c) -> hov 1 (str"[" ++ prlist_with_sep spc Id.print l ++ str"]") - ++ str "," ++ spc () ++ Termops.print_constr (EConstr.Unsafe.to_constr c)) + ++ str "," ++ spc () ++ Termops.print_constr c) let ppconstrunderbindersidmap l = pp (prconstrunderbindersidmap l) @@ -159,7 +159,7 @@ let prdelta s = pp (Mod_subst.debug_pr_delta s) let pp_idpred s = pp (pr_idpred s) let pp_cpred s = pp (pr_cpred s) let pp_transparent_state s = pp (pr_transparent_state s) -let pp_stack_t n = pp (Reductionops.Stack.pr Termops.print_constr n) +let pp_stack_t n = pp (Reductionops.Stack.pr (EConstr.of_constr %> Termops.print_constr) n) let pp_cst_stack_t n = pp (Reductionops.Cst_stack.pr n) let pp_state_t n = pp (Reductionops.pr_state n) diff --git a/engine/evd.ml b/engine/evd.ml index d8a658e3e..9c05c6c02 100644 --- a/engine/evd.ml +++ b/engine/evd.ml @@ -1265,7 +1265,7 @@ let protect f x = let (f_print_constr, print_constr_hook) = Hook.make () -let print_constr a = protect (fun c -> Hook.get f_print_constr (Global.env ()) c) a +let print_constr a = protect (fun c -> Hook.get f_print_constr (Global.env ()) empty c) a let pr_meta_map mmap = let pr_name = function @@ -1423,11 +1423,11 @@ let pr_evar_constraints pbs = Namegen.make_all_name_different env in print_env_short env ++ spc () ++ str "|-" ++ spc () ++ - Hook.get f_print_constr env t1 ++ spc () ++ + Hook.get f_print_constr env empty t1 ++ spc () ++ str (match pbty with | Reduction.CONV -> "==" | Reduction.CUMUL -> "<=") ++ - spc () ++ Hook.get f_print_constr env t2 + spc () ++ Hook.get f_print_constr env empty t2 in prlist_with_sep fnl pr_evconstr pbs diff --git a/engine/evd.mli b/engine/evd.mli index 5c9effd4b..4e8284675 100644 --- a/engine/evd.mli +++ b/engine/evd.mli @@ -607,7 +607,7 @@ val pr_evar_suggested_name : existential_key -> evar_map -> Id.t (** {5 Debug pretty-printers} *) -val print_constr_hook : (Environ.env -> constr -> Pp.std_ppcmds) Hook.t +val print_constr_hook : (Environ.env -> evar_map -> constr -> Pp.std_ppcmds) Hook.t 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 diff --git a/engine/termops.ml b/engine/termops.ml index 465832c10..46e9ad927 100644 --- a/engine/termops.ml +++ b/engine/termops.ml @@ -97,11 +97,11 @@ let rec pr_constr c = match kind_of_term c with 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 -let () = Hook.set Evd.print_constr_hook (fun env c -> !term_printer env c) +let () = Hook.set Evd.print_constr_hook (fun env sigma c -> !term_printer env sigma (EConstr.of_constr c)) let pr_var_decl env decl = let open NamedDecl in @@ -109,9 +109,10 @@ let pr_var_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" := " ++ 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)) @@ -121,9 +122,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) diff --git a/engine/termops.mli b/engine/termops.mli index d7d9c743d..3f924cfa1 100644 --- a/engine/termops.mli +++ b/engine/termops.mli @@ -21,9 +21,9 @@ val pr_sort_family : sorts_family -> std_ppcmds val pr_fix : ('a -> std_ppcmds) -> ('a, 'a) pfixpoint -> std_ppcmds (** debug printer: do not use to display terms to the casual user... *) -val set_print_constr : (env -> Constr.constr -> std_ppcmds) -> unit -val print_constr : Constr.constr -> std_ppcmds -val print_constr_env : env -> Constr.constr -> std_ppcmds +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 diff --git a/interp/stdarg.mli b/interp/stdarg.mli index 3047d2bce..113fe40ba 100644 --- a/interp/stdarg.mli +++ b/interp/stdarg.mli @@ -11,6 +11,7 @@ open Loc open Names open Term +open EConstr open Libnames open Globnames open Genredexpr @@ -57,12 +58,12 @@ val wit_open_constr : val wit_constr_with_bindings : (constr_expr with_bindings, glob_constr_and_expr with_bindings, - EConstr.constr with_bindings delayed_open) genarg_type + constr with_bindings delayed_open) genarg_type val wit_bindings : (constr_expr bindings, glob_constr_and_expr bindings, - EConstr.constr bindings delayed_open) genarg_type + constr bindings delayed_open) genarg_type val wit_red_expr : ((constr_expr,reference or_by_notation,constr_expr) red_expr_gen, diff --git a/ltac/coretactics.ml4 b/ltac/coretactics.ml4 index 20d9640fc..28ff6df83 100644 --- a/ltac/coretactics.ml4 +++ b/ltac/coretactics.ml4 @@ -27,7 +27,7 @@ TACTIC EXTEND reflexivity END TACTIC EXTEND exact - [ "exact" casted_constr(c) ] -> [ Tactics.exact_no_check (EConstr.of_constr c) ] + [ "exact" casted_constr(c) ] -> [ Tactics.exact_no_check c ] END TACTIC EXTEND assumption @@ -39,35 +39,35 @@ TACTIC EXTEND etransitivity END TACTIC EXTEND cut - [ "cut" constr(c) ] -> [ Tactics.cut (EConstr.of_constr c) ] + [ "cut" constr(c) ] -> [ Tactics.cut c ] END TACTIC EXTEND exact_no_check - [ "exact_no_check" constr(c) ] -> [ Tactics.exact_no_check (EConstr.of_constr c) ] + [ "exact_no_check" constr(c) ] -> [ Tactics.exact_no_check c ] END TACTIC EXTEND vm_cast_no_check - [ "vm_cast_no_check" constr(c) ] -> [ Tactics.vm_cast_no_check (EConstr.of_constr c) ] + [ "vm_cast_no_check" constr(c) ] -> [ Tactics.vm_cast_no_check c ] END TACTIC EXTEND native_cast_no_check - [ "native_cast_no_check" constr(c) ] -> [ Tactics.native_cast_no_check (EConstr.of_constr c) ] + [ "native_cast_no_check" constr(c) ] -> [ Tactics.native_cast_no_check c ] END TACTIC EXTEND casetype - [ "casetype" constr(c) ] -> [ Tactics.case_type (EConstr.of_constr c) ] + [ "casetype" constr(c) ] -> [ Tactics.case_type c ] END TACTIC EXTEND elimtype - [ "elimtype" constr(c) ] -> [ Tactics.elim_type (EConstr.of_constr c) ] + [ "elimtype" constr(c) ] -> [ Tactics.elim_type c ] END TACTIC EXTEND lapply - [ "lapply" constr(c) ] -> [ Tactics.cut_and_apply (EConstr.of_constr c) ] + [ "lapply" constr(c) ] -> [ Tactics.cut_and_apply c ] END TACTIC EXTEND transitivity - [ "transitivity" constr(c) ] -> [ Tactics.intros_transitivity (Some (EConstr.of_constr c)) ] + [ "transitivity" constr(c) ] -> [ Tactics.intros_transitivity (Some c) ] END (** Left *) @@ -297,7 +297,7 @@ END (* Generalize dependent *) TACTIC EXTEND generalize_dependent - [ "generalize" "dependent" constr(c) ] -> [ Tactics.generalize_dep (EConstr.of_constr c) ] + [ "generalize" "dependent" constr(c) ] -> [ Tactics.generalize_dep c ] END (* Table of "pervasives" macros tactics (e.g. auto, simpl, etc.) *) diff --git a/ltac/evar_tactics.ml b/ltac/evar_tactics.ml index 99023fdbb..6b94da28a 100644 --- a/ltac/evar_tactics.ml +++ b/ltac/evar_tactics.ml @@ -77,15 +77,16 @@ let let_evar name typ = let sigma = Tacmach.New.project gl in let env = Proofview.Goal.env gl in let sigma = ref sigma in - let _ = Typing.e_sort_of env sigma (EConstr.of_constr typ) in + let _ = Typing.e_sort_of env sigma typ in let sigma = Sigma.Unsafe.of_evar_map !sigma in let id = match name with | Names.Anonymous -> + let typ = EConstr.Unsafe.to_constr typ in let id = Namegen.id_of_name_using_hdchar env typ name in Namegen.next_ident_away_in_goal id (Termops.ids_of_named_context (Environ.named_context env)) | Names.Name id -> id in - let Sigma (evar, sigma, p) = Evarutil.new_evar env sigma ~src ~naming:(Misctypes.IntroFresh id) (EConstr.of_constr typ) in + let Sigma (evar, sigma, p) = Evarutil.new_evar env sigma ~src ~naming:(Misctypes.IntroFresh id) typ in let tac = (Tactics.letin_tac None (Names.Name id) evar None Locusops.nowhere) in diff --git a/ltac/evar_tactics.mli b/ltac/evar_tactics.mli index e67540c05..41266e61c 100644 --- a/ltac/evar_tactics.mli +++ b/ltac/evar_tactics.mli @@ -16,4 +16,4 @@ val instantiate_tac : int -> Tacinterp.interp_sign * Glob_term.glob_constr -> val instantiate_tac_by_name : Id.t -> Tacinterp.interp_sign * Glob_term.glob_constr -> unit Proofview.tactic -val let_evar : Name.t -> Term.types -> unit Proofview.tactic +val let_evar : Name.t -> EConstr.types -> unit Proofview.tactic diff --git a/ltac/extraargs.mli b/ltac/extraargs.mli index b12187e18..7d4bccfad 100644 --- a/ltac/extraargs.mli +++ b/ltac/extraargs.mli @@ -38,12 +38,12 @@ val wit_lglob : val wit_lconstr : (constr_expr, Tacexpr.glob_constr_and_expr, - Constr.t) Genarg.genarg_type + EConstr.t) Genarg.genarg_type val wit_casted_constr : (constr_expr, Tacexpr.glob_constr_and_expr, - Constr.t) Genarg.genarg_type + EConstr.t) Genarg.genarg_type val glob : constr_expr Pcoq.Gram.entry val lglob : constr_expr Pcoq.Gram.entry diff --git a/ltac/extratactics.ml4 b/ltac/extratactics.ml4 index bf8481b52..65c75703b 100644 --- a/ltac/extratactics.ml4 +++ b/ltac/extratactics.ml4 @@ -47,16 +47,16 @@ let with_delayed_uconstr ist c tac = let replace_in_clause_maybe_by ist c1 c2 cl tac = with_delayed_uconstr ist c1 - (fun c1 -> replace_in_clause_maybe_by (EConstr.of_constr c1) c2 cl (Option.map (Tacinterp.tactic_of_value ist) tac)) + (fun c1 -> replace_in_clause_maybe_by c1 c2 cl (Option.map (Tacinterp.tactic_of_value ist) tac)) let replace_term ist dir_opt c cl = - with_delayed_uconstr ist c (fun c -> replace_term dir_opt (EConstr.of_constr c) cl) + with_delayed_uconstr ist c (fun c -> replace_term dir_opt c cl) let clause = Pltac.clause_dft_concl TACTIC EXTEND replace ["replace" uconstr(c1) "with" constr(c2) clause(cl) by_arg_tac(tac) ] --> [ replace_in_clause_maybe_by ist c1 (EConstr.of_constr c2) cl tac ] +-> [ replace_in_clause_maybe_by ist c1 c2 cl tac ] END TACTIC EXTEND replace_term_left @@ -153,9 +153,9 @@ let injHyp id = injection_main false { delayed = fun env sigma -> Sigma.here (EConstr.mkVar id, NoBindings) sigma } TACTIC EXTEND dependent_rewrite -| [ "dependent" "rewrite" orient(b) constr(c) ] -> [ rewriteInConcl b (EConstr.of_constr c) ] +| [ "dependent" "rewrite" orient(b) constr(c) ] -> [ rewriteInConcl b c ] | [ "dependent" "rewrite" orient(b) constr(c) "in" hyp(id) ] - -> [ rewriteInHyp b (EConstr.of_constr c) id ] + -> [ rewriteInHyp b c id ] END (** To be deprecated?, "cutrewrite (t=u) as <-" is equivalent to @@ -163,20 +163,20 @@ END "cutrewrite (t=u) as ->" is equivalent to "enough (t=u) as ->". *) TACTIC EXTEND cut_rewrite -| [ "cutrewrite" orient(b) constr(eqn) ] -> [ cutRewriteInConcl b (EConstr.of_constr eqn) ] +| [ "cutrewrite" orient(b) constr(eqn) ] -> [ cutRewriteInConcl b eqn ] | [ "cutrewrite" orient(b) constr(eqn) "in" hyp(id) ] - -> [ cutRewriteInHyp b (EConstr.of_constr eqn) id ] + -> [ cutRewriteInHyp b eqn id ] END (**********************************************************************) (* Decompose *) TACTIC EXTEND decompose_sum -| [ "decompose" "sum" constr(c) ] -> [ Elim.h_decompose_or (EConstr.of_constr c) ] +| [ "decompose" "sum" constr(c) ] -> [ Elim.h_decompose_or c ] END TACTIC EXTEND decompose_record -| [ "decompose" "record" constr(c) ] -> [ Elim.h_decompose_and (EConstr.of_constr c) ] +| [ "decompose" "record" constr(c) ] -> [ Elim.h_decompose_and c ] END (**********************************************************************) @@ -185,7 +185,7 @@ END open Contradiction TACTIC EXTEND absurd - [ "absurd" constr(c) ] -> [ absurd (EConstr.of_constr c) ] + [ "absurd" constr(c) ] -> [ absurd c ] END let onSomeWithHoles tac = function @@ -235,7 +235,7 @@ END let rewrite_star ist clause orient occs c (tac : Geninterp.Val.t option) = let tac' = Option.map (fun t -> Tacinterp.tactic_of_value ist t, FirstSolved) tac in with_delayed_uconstr ist c - (fun c -> general_rewrite_ebindings_clause clause orient occs ?tac:tac' true true (EConstr.of_constr c,NoBindings) true) + (fun c -> general_rewrite_ebindings_clause clause orient occs ?tac:tac' true true (c,NoBindings) true) TACTIC EXTEND rewrite_star | [ "rewrite" "*" orient(o) uconstr(c) "in" hyp(id) "at" occurrences(occ) by_arg_tac(tac) ] -> @@ -362,7 +362,7 @@ let refine_tac ist simple c = let c = Pretyping.type_uconstr ~flags ~expected_type ist c in let update = { run = fun sigma -> let Sigma (c, sigma, p) = c.delayed env sigma in - Sigma (EConstr.of_constr c, sigma, p) + Sigma (c, sigma, p) } in let refine = Refine.refine ~unsafe:true update in if simple then refine @@ -516,13 +516,13 @@ let add_transitivity_lemma left lem = (* Vernacular syntax *) TACTIC EXTEND stepl -| ["stepl" constr(c) "by" tactic(tac) ] -> [ step true (EConstr.of_constr c) (Tacinterp.tactic_of_value ist tac) ] -| ["stepl" constr(c) ] -> [ step true (EConstr.of_constr c) (Proofview.tclUNIT ()) ] +| ["stepl" constr(c) "by" tactic(tac) ] -> [ step true c (Tacinterp.tactic_of_value ist tac) ] +| ["stepl" constr(c) ] -> [ step true c (Proofview.tclUNIT ()) ] END TACTIC EXTEND stepr -| ["stepr" constr(c) "by" tactic(tac) ] -> [ step false (EConstr.of_constr c) (Tacinterp.tactic_of_value ist tac) ] -| ["stepr" constr(c) ] -> [ step false (EConstr.of_constr c) (Proofview.tclUNIT ()) ] +| ["stepr" constr(c) "by" tactic(tac) ] -> [ step false c (Tacinterp.tactic_of_value ist tac) ] +| ["stepr" constr(c) ] -> [ step false c (Proofview.tclUNIT ()) ] END VERNAC COMMAND EXTEND AddStepl CLASSIFIED AS SIDEFF @@ -645,6 +645,8 @@ let subst_hole_with_term occ tc t = open Tacmach let hResolve id c occ t = + let c = EConstr.Unsafe.to_constr c in + let t = EConstr.Unsafe.to_constr t in Proofview.Goal.nf_s_enter { s_enter = begin fun gl -> let sigma = Proofview.Goal.sigma gl in let sigma = Sigma.to_evar_map sigma in @@ -819,8 +821,9 @@ END let eq_constr x y = Proofview.Goal.enter { enter = begin fun gl -> let evd = Tacmach.New.project gl in - if Evarutil.eq_constr_univs_test evd evd x y then Proofview.tclUNIT () - else Tacticals.New.tclFAIL 0 (str "Not equal") + match EConstr.eq_constr_universes evd x y with + | Some _ -> Proofview.tclUNIT () + | None -> Tacticals.New.tclFAIL 0 (str "Not equal") end } TACTIC EXTEND constr_eq @@ -830,15 +833,13 @@ END TACTIC EXTEND constr_eq_nounivs | [ "constr_eq_nounivs" constr(x) constr(y) ] -> [ Proofview.tclEVARMAP >>= fun sigma -> - let x = EConstr.of_constr x in - let y = EConstr.of_constr y in if eq_constr_nounivs sigma x y then Proofview.tclUNIT () else Tacticals.New.tclFAIL 0 (str "Not equal") ] END TACTIC EXTEND is_evar | [ "is_evar" constr(x) ] -> [ Proofview.tclEVARMAP >>= fun sigma -> - match EConstr.kind sigma (EConstr.of_constr x) with + match EConstr.kind sigma x with | Evar _ -> Proofview.tclUNIT () | _ -> Tacticals.New.tclFAIL 0 (str "Not an evar") ] @@ -871,14 +872,14 @@ has_evar c TACTIC EXTEND has_evar | [ "has_evar" constr(x) ] -> [ Proofview.tclEVARMAP >>= fun sigma -> - if has_evar sigma (EConstr.of_constr x) then Proofview.tclUNIT () else Tacticals.New.tclFAIL 0 (str "No evars") + if has_evar sigma x then Proofview.tclUNIT () else Tacticals.New.tclFAIL 0 (str "No evars") ] END TACTIC EXTEND is_hyp | [ "is_var" constr(x) ] -> [ Proofview.tclEVARMAP >>= fun sigma -> - match EConstr.kind sigma (EConstr.of_constr x) with + match EConstr.kind sigma x with | Var _ -> Proofview.tclUNIT () | _ -> Tacticals.New.tclFAIL 0 (str "Not a variable or hypothesis") ] END @@ -886,7 +887,7 @@ END TACTIC EXTEND is_fix | [ "is_fix" constr(x) ] -> [ Proofview.tclEVARMAP >>= fun sigma -> - match EConstr.kind sigma (EConstr.of_constr x) with + match EConstr.kind sigma x with | Fix _ -> Proofview.tclUNIT () | _ -> Tacticals.New.tclFAIL 0 (Pp.str "not a fix definition") ] END;; @@ -894,7 +895,7 @@ END;; TACTIC EXTEND is_cofix | [ "is_cofix" constr(x) ] -> [ Proofview.tclEVARMAP >>= fun sigma -> - match EConstr.kind sigma (EConstr.of_constr x) with + match EConstr.kind sigma x with | CoFix _ -> Proofview.tclUNIT () | _ -> Tacticals.New.tclFAIL 0 (Pp.str "not a cofix definition") ] END;; @@ -902,7 +903,7 @@ END;; TACTIC EXTEND is_ind | [ "is_ind" constr(x) ] -> [ Proofview.tclEVARMAP >>= fun sigma -> - match EConstr.kind sigma (EConstr.of_constr x) with + match EConstr.kind sigma x with | Ind _ -> Proofview.tclUNIT () | _ -> Tacticals.New.tclFAIL 0 (Pp.str "not an (co)inductive datatype") ] END;; @@ -910,7 +911,7 @@ END;; TACTIC EXTEND is_constructor | [ "is_constructor" constr(x) ] -> [ Proofview.tclEVARMAP >>= fun sigma -> - match EConstr.kind sigma (EConstr.of_constr x) with + match EConstr.kind sigma x with | Construct _ -> Proofview.tclUNIT () | _ -> Tacticals.New.tclFAIL 0 (Pp.str "not a constructor") ] END;; @@ -918,7 +919,7 @@ END;; TACTIC EXTEND is_proj | [ "is_proj" constr(x) ] -> [ Proofview.tclEVARMAP >>= fun sigma -> - match EConstr.kind sigma (EConstr.of_constr x) with + match EConstr.kind sigma x with | Proj _ -> Proofview.tclUNIT () | _ -> Tacticals.New.tclFAIL 0 (Pp.str "not a primitive projection") ] END;; @@ -926,7 +927,7 @@ END;; TACTIC EXTEND is_const | [ "is_const" constr(x) ] -> [ Proofview.tclEVARMAP >>= fun sigma -> - match EConstr.kind sigma (EConstr.of_constr x) with + match EConstr.kind sigma x with | Const _ -> Proofview.tclUNIT () | _ -> Tacticals.New.tclFAIL 0 (Pp.str "not a constant") ] END;; @@ -1080,7 +1081,7 @@ let decompose l c = end } TACTIC EXTEND decompose -| [ "decompose" "[" ne_constr_list(l) "]" constr(c) ] -> [ decompose (List.map EConstr.of_constr l) (EConstr.of_constr c) ] +| [ "decompose" "[" ne_constr_list(l) "]" constr(c) ] -> [ decompose l c ] END (** library/keys *) diff --git a/ltac/g_auto.ml4 b/ltac/g_auto.ml4 index b5e56d93b..d4d3b9d66 100644 --- a/ltac/g_auto.ml4 +++ b/ltac/g_auto.ml4 @@ -28,7 +28,7 @@ TACTIC EXTEND eassumption END TACTIC EXTEND eexact -| [ "eexact" constr(c) ] -> [ Eauto.e_give_exact (EConstr.of_constr c) ] +| [ "eexact" constr(c) ] -> [ Eauto.e_give_exact c ] END let pr_hintbases _prc _prlc _prt = Pptactic.pr_hintbases @@ -51,7 +51,7 @@ let eval_uconstrs ist cs = } in let map c = { delayed = fun env sigma -> let Sigma.Sigma (c, sigma, p) = c.delayed env sigma in - Sigma.Sigma (EConstr.of_constr c, sigma, p) + Sigma.Sigma (c, sigma, p) } in List.map (fun c -> map (Pretyping.type_uconstr ~flags ist c)) cs @@ -151,7 +151,7 @@ TACTIC EXTEND autounfold_one TACTIC EXTEND autounfoldify | [ "autounfoldify" constr(x) ] -> [ Proofview.tclEVARMAP >>= fun sigma -> - let db = match EConstr.kind sigma (EConstr.of_constr x) with + let db = match EConstr.kind sigma x with | Term.Const (c,_) -> Names.Label.to_string (Names.con_label c) | _ -> assert false in Eauto.autounfold ["core";db] Locusops.onConcl @@ -159,7 +159,7 @@ TACTIC EXTEND autounfoldify END TACTIC EXTEND unify -| ["unify" constr(x) constr(y) ] -> [ Tactics.unify (EConstr.of_constr x) (EConstr.of_constr y) ] +| ["unify" constr(x) constr(y) ] -> [ Tactics.unify x y ] | ["unify" constr(x) constr(y) "with" preident(base) ] -> [ let table = try Some (Hints.searchtable_map base) with Not_found -> None in match table with @@ -168,13 +168,13 @@ TACTIC EXTEND unify Tacticals.New.tclZEROMSG msg | Some t -> let state = Hints.Hint_db.transparent_state t in - Tactics.unify ~state (EConstr.of_constr x) (EConstr.of_constr y) + Tactics.unify ~state x y ] END TACTIC EXTEND convert_concl_no_check -| ["convert_concl_no_check" constr(x) ] -> [ Tactics.convert_concl_no_check (EConstr.of_constr x) Term.DEFAULTcast ] +| ["convert_concl_no_check" constr(x) ] -> [ Tactics.convert_concl_no_check x Term.DEFAULTcast ] END let pr_hints_path_atom _ _ _ = Hints.pp_hints_path_atom diff --git a/ltac/g_class.ml4 b/ltac/g_class.ml4 index a983a4fea..b7f0881f1 100644 --- a/ltac/g_class.ml4 +++ b/ltac/g_class.ml4 @@ -62,19 +62,19 @@ TACTIC EXTEND typeclasses_eauto END TACTIC EXTEND head_of_constr - [ "head_of_constr" ident(h) constr(c) ] -> [ head_of_constr h (EConstr.of_constr c) ] + [ "head_of_constr" ident(h) constr(c) ] -> [ head_of_constr h c ] END TACTIC EXTEND not_evar - [ "not_evar" constr(ty) ] -> [ not_evar (EConstr.of_constr ty) ] + [ "not_evar" constr(ty) ] -> [ not_evar ty ] END TACTIC EXTEND is_ground - [ "is_ground" constr(ty) ] -> [ Proofview.V82.tactic (is_ground (EConstr.of_constr ty)) ] + [ "is_ground" constr(ty) ] -> [ Proofview.V82.tactic (is_ground ty) ] END TACTIC EXTEND autoapply - [ "autoapply" constr(c) "using" preident(i) ] -> [ Proofview.V82.tactic (autoapply (EConstr.of_constr c) i) ] + [ "autoapply" constr(c) "using" preident(i) ] -> [ Proofview.V82.tactic (autoapply c i) ] END (** TODO: DEPRECATE *) diff --git a/ltac/g_eqdecide.ml4 b/ltac/g_eqdecide.ml4 index 6a49ea830..905653281 100644 --- a/ltac/g_eqdecide.ml4 +++ b/ltac/g_eqdecide.ml4 @@ -23,5 +23,5 @@ TACTIC EXTEND decide_equality END TACTIC EXTEND compare -| [ "compare" constr(c1) constr(c2) ] -> [ compare (EConstr.of_constr c1) (EConstr.of_constr c2) ] +| [ "compare" constr(c1) constr(c2) ] -> [ compare c1 c2 ] END diff --git a/ltac/g_rewrite.ml4 b/ltac/g_rewrite.ml4 index bae5a516c..b1c4f58eb 100644 --- a/ltac/g_rewrite.ml4 +++ b/ltac/g_rewrite.ml4 @@ -265,7 +265,7 @@ TACTIC EXTEND setoid_reflexivity END TACTIC EXTEND setoid_transitivity - [ "setoid_transitivity" constr(t) ] -> [ setoid_transitivity (Some (EConstr.of_constr t)) ] + [ "setoid_transitivity" constr(t) ] -> [ setoid_transitivity (Some t) ] | [ "setoid_etransitivity" ] -> [ setoid_transitivity None ] END diff --git a/ltac/pptactic.ml b/ltac/pptactic.ml index 934830f4d..b7dd37cdd 100644 --- a/ltac/pptactic.ml +++ b/ltac/pptactic.ml @@ -59,8 +59,8 @@ type 'a glob_extra_genarg_printer = 'a -> std_ppcmds type 'a extra_genarg_printer = - (Term.constr -> std_ppcmds) -> - (Term.constr -> std_ppcmds) -> + (EConstr.constr -> std_ppcmds) -> + (EConstr.constr -> std_ppcmds) -> (tolerability -> Val.t -> std_ppcmds) -> 'a -> std_ppcmds @@ -1167,15 +1167,15 @@ module Make | _ -> error "Cannot translate fix tactic: not enough products" in strip_ty [] n ty - let pr_atomic_tactic_level env n t = + let pr_atomic_tactic_level env sigma n t = let prtac n (t:atomic_tactic_expr) = let pr = { pr_tactic = (fun _ _ -> str ""); - pr_constr = (fun c -> pr_constr_env env Evd.empty (EConstr.Unsafe.to_constr c)); + pr_constr = (fun c -> pr_econstr_env env sigma c); pr_dconstr = pr_and_constr_expr (pr_glob_constr_env env); - pr_lconstr = (fun c -> pr_lconstr_env env Evd.empty (EConstr.Unsafe.to_constr c)); - pr_pattern = pr_constr_pattern_env env Evd.empty; - pr_lpattern = pr_lconstr_pattern_env env Evd.empty; + pr_lconstr = (fun c -> pr_leconstr_env env sigma c); + pr_pattern = pr_constr_pattern_env env sigma; + pr_lpattern = pr_lconstr_pattern_env env sigma; pr_constant = pr_evaluable_reference_env env; pr_reference = pr_located pr_ltac_constant; pr_name = pr_id; @@ -1206,7 +1206,7 @@ module Make let pr_extend pr lev ml args = pr_extend_gen pr lev ml args - let pr_atomic_tactic env = pr_atomic_tactic_level env ltop + let pr_atomic_tactic env sigma c = pr_atomic_tactic_level env sigma ltop c end @@ -1255,7 +1255,7 @@ let declare_extra_genarg_pprule wit in let h x = let env = Global.env () in - h (pr_constr_env env Evd.empty) (pr_lconstr_env env Evd.empty) (fun _ _ -> str "") x + h (pr_econstr_env env Evd.empty) (pr_leconstr_env env Evd.empty) (fun _ _ -> str "") x in Genprint.register_print0 wit f g h @@ -1285,7 +1285,7 @@ let () = wit_intro_pattern (Miscprint.pr_intro_pattern pr_constr_expr) (Miscprint.pr_intro_pattern (fun (c,_) -> pr_glob_constr c)) - (Miscprint.pr_intro_pattern (fun c -> pr_constr (EConstr.Unsafe.to_constr (fst (run_delayed c))))); + (Miscprint.pr_intro_pattern (fun c -> pr_econstr (fst (run_delayed c)))); Genprint.register_print0 wit_clause_dft_concl (pr_clauses (Some true) pr_lident) @@ -1296,7 +1296,7 @@ let () = wit_constr Ppconstr.pr_constr_expr (fun (c, _) -> Printer.pr_glob_constr c) - Printer.pr_constr + Printer.pr_econstr ; Genprint.register_print0 wit_uconstr @@ -1308,25 +1308,25 @@ let () = wit_open_constr Ppconstr.pr_constr_expr (fun (c, _) -> Printer.pr_glob_constr c) - Printer.pr_constr + Printer.pr_econstr ; Genprint.register_print0 wit_red_expr (pr_red_expr (pr_constr_expr, pr_lconstr_expr, pr_or_by_notation pr_reference, pr_constr_pattern_expr)) (pr_red_expr (pr_and_constr_expr pr_glob_constr, pr_and_constr_expr pr_lglob_constr, pr_or_var (pr_and_short_name pr_evaluable_reference), pr_pat_and_constr_expr pr_glob_constr)) - (pr_red_expr (pr_constr, pr_lconstr, pr_evaluable_reference, pr_constr_pattern)); + (pr_red_expr (pr_econstr, pr_leconstr, pr_evaluable_reference, pr_constr_pattern)); Genprint.register_print0 wit_quant_hyp pr_quantified_hypothesis pr_quantified_hypothesis pr_quantified_hypothesis; Genprint.register_print0 wit_bindings (pr_bindings_no_with pr_constr_expr pr_lconstr_expr) (pr_bindings_no_with (pr_and_constr_expr pr_glob_constr) (pr_and_constr_expr pr_lglob_constr)) - (fun it -> pr_bindings_no_with (EConstr.Unsafe.to_constr %> pr_constr) (EConstr.Unsafe.to_constr %> pr_lconstr) (fst (run_delayed it))); + (fun it -> pr_bindings_no_with pr_econstr pr_leconstr (fst (run_delayed it))); Genprint.register_print0 wit_constr_with_bindings (pr_with_bindings pr_constr_expr pr_lconstr_expr) (pr_with_bindings (pr_and_constr_expr pr_glob_constr) (pr_and_constr_expr pr_lglob_constr)) - (fun it -> pr_with_bindings (EConstr.Unsafe.to_constr %> pr_constr) (EConstr.Unsafe.to_constr %> pr_lconstr) (fst (run_delayed it))); + (fun it -> pr_with_bindings pr_econstr pr_leconstr (fst (run_delayed it))); Genprint.register_print0 Tacarg.wit_destruction_arg (pr_destruction_arg pr_constr_expr pr_lconstr_expr) (pr_destruction_arg (pr_and_constr_expr pr_glob_constr) (pr_and_constr_expr pr_lglob_constr)) - (fun it -> pr_destruction_arg (EConstr.Unsafe.to_constr %> pr_constr) (EConstr.Unsafe.to_constr %> pr_lconstr) (run_delayed_destruction_arg it)); + (fun it -> pr_destruction_arg pr_econstr pr_leconstr (run_delayed_destruction_arg it)); Genprint.register_print0 Stdarg.wit_int int int int; Genprint.register_print0 Stdarg.wit_bool pr_bool pr_bool pr_bool; Genprint.register_print0 Stdarg.wit_unit pr_unit pr_unit pr_unit; diff --git a/ltac/pptactic.mli b/ltac/pptactic.mli index 86e3ea548..5fff3062d 100644 --- a/ltac/pptactic.mli +++ b/ltac/pptactic.mli @@ -34,8 +34,8 @@ type 'a glob_extra_genarg_printer = 'a -> std_ppcmds type 'a extra_genarg_printer = - (Term.constr -> std_ppcmds) -> - (Term.constr -> std_ppcmds) -> + (EConstr.constr -> std_ppcmds) -> + (EConstr.constr -> std_ppcmds) -> (tolerability -> Val.t -> std_ppcmds) -> 'a -> std_ppcmds diff --git a/ltac/pptacticsig.mli b/ltac/pptacticsig.mli index 74ddd377a..dfbc3b3ed 100644 --- a/ltac/pptacticsig.mli +++ b/ltac/pptacticsig.mli @@ -61,7 +61,7 @@ module type Pp = sig val pr_glob_tactic : env -> glob_tactic_expr -> std_ppcmds - val pr_atomic_tactic : env -> atomic_tactic_expr -> std_ppcmds + val pr_atomic_tactic : env -> Evd.evar_map -> atomic_tactic_expr -> std_ppcmds val pr_hintbases : string list option -> std_ppcmds diff --git a/ltac/taccoerce.ml b/ltac/taccoerce.ml index 288e12d0b..114b8dda0 100644 --- a/ltac/taccoerce.ml +++ b/ltac/taccoerce.ml @@ -9,6 +9,7 @@ open Util open Names open Term +open EConstr open Pattern open Misctypes open Genarg @@ -64,7 +65,7 @@ let to_constr v = Some c else if has_type v (topwit wit_constr_under_binders) then let vars, c = out_gen (topwit wit_constr_under_binders) v in - match vars with [] -> Some (EConstr.Unsafe.to_constr c) | _ -> None + match vars with [] -> Some c | _ -> None else None let of_uconstr c = in_gen (topwit wit_uconstr) c @@ -106,7 +107,7 @@ let coerce_to_constr_context v = else raise (CannotCoerceTo "a term context") (* Interprets an identifier which must be fresh *) -let coerce_var_to_ident fresh env v = +let coerce_var_to_ident fresh env sigma v = let v = Value.normalize v in let fail () = raise (CannotCoerceTo "a fresh identifier") in if has_type v (topwit wit_intro_pattern) then @@ -119,15 +120,16 @@ let coerce_var_to_ident fresh env v = | None -> fail () | Some c -> (* We need it fresh for intro e.g. in "Tac H = clear H; intro H" *) - if isVar c && not (fresh && is_variable env (destVar c)) then - destVar c + if isVar sigma c && not (fresh && is_variable env (destVar sigma c)) then + destVar sigma c else fail () (* Interprets, if possible, a constr to an identifier which may not be fresh but suitable to be given to the fresh tactic. Works for vars, constants, inductive, constructors and sorts. *) -let coerce_to_ident_not_fresh g env v = +let coerce_to_ident_not_fresh env sigma v = +let g = sigma in let id_of_name = function | Names.Anonymous -> Id.of_string "x" | Names.Name x -> x in @@ -143,7 +145,7 @@ let id_of_name = function match Value.to_constr v with | None -> fail () | Some c -> - match Constr.kind c with + match EConstr.kind sigma c with | Var id -> id | Meta m -> id_of_name (Evd.meta_name g m) | Evar (kn,_) -> @@ -169,7 +171,7 @@ let id_of_name = function | _ -> fail() -let coerce_to_intro_pattern env v = +let coerce_to_intro_pattern env sigma v = let v = Value.normalize v in if has_type v (topwit wit_intro_pattern) then snd (out_gen (topwit wit_intro_pattern) v) @@ -177,14 +179,14 @@ let coerce_to_intro_pattern env v = let id = out_gen (topwit wit_var) v in IntroNaming (IntroIdentifier id) else match Value.to_constr v with - | Some c when isVar c -> + | Some c when isVar sigma c -> (* This happens e.g. in definitions like "Tac H = clear H; intro H" *) (* but also in "destruct H as (H,H')" *) - IntroNaming (IntroIdentifier (destVar c)) + IntroNaming (IntroIdentifier (destVar sigma c)) | _ -> raise (CannotCoerceTo "an introduction pattern") -let coerce_to_intro_pattern_naming env v = - match coerce_to_intro_pattern env v with +let coerce_to_intro_pattern_naming env sigma v = + match coerce_to_intro_pattern env sigma v with | IntroNaming pat -> pat | _ -> raise (CannotCoerceTo "a naming introduction pattern") @@ -212,7 +214,7 @@ let coerce_to_constr env v = | _ -> fail () else if has_type v (topwit wit_constr) then let c = out_gen (topwit wit_constr) v in - ([], EConstr.of_constr c) + ([], c) else if has_type v (topwit wit_constr_under_binders) then out_gen (topwit wit_constr_under_binders) v else if has_type v (topwit wit_var) then @@ -229,11 +231,10 @@ let coerce_to_uconstr env v = let coerce_to_closed_constr env v = let ids,c = coerce_to_constr env v in - let c = EConstr.Unsafe.to_constr c in let () = if not (List.is_empty ids) then raise (CannotCoerceTo "a term") in c -let coerce_to_evaluable_ref env v = +let coerce_to_evaluable_ref env sigma v = let fail () = raise (CannotCoerceTo "an evaluable reference") in let v = Value.normalize v in let ev = @@ -254,8 +255,8 @@ let coerce_to_evaluable_ref env v = | IndRef _ | ConstructRef _ -> fail () else match Value.to_constr v with - | Some c when isConst c -> EvalConstRef (Univ.out_punivs (destConst c)) - | Some c when isVar c -> EvalVarRef (destVar c) + | Some c when isConst sigma c -> EvalConstRef (Univ.out_punivs (destConst sigma c)) + | Some c when isVar sigma c -> EvalVarRef (destVar sigma c) | _ -> fail () in if Tacred.is_evaluable env ev then ev else fail () @@ -267,14 +268,14 @@ let coerce_to_constr_list env v = List.map map l | None -> raise (CannotCoerceTo "a term list") -let coerce_to_intro_pattern_list loc env v = +let coerce_to_intro_pattern_list loc env sigma v = match Value.to_list v with | None -> raise (CannotCoerceTo "an intro pattern list") | Some l -> - let map v = (loc, coerce_to_intro_pattern env v) in + let map v = (loc, coerce_to_intro_pattern env sigma v) in List.map map l -let coerce_to_hyp env v = +let coerce_to_hyp env sigma v = let fail () = raise (CannotCoerceTo "a variable") in let v = Value.normalize v in if has_type v (topwit wit_intro_pattern) then @@ -285,31 +286,31 @@ let coerce_to_hyp env v = let id = out_gen (topwit wit_var) v in if is_variable env id then id else fail () else match Value.to_constr v with - | Some c when isVar c -> destVar c + | Some c when isVar sigma c -> destVar sigma c | _ -> fail () -let coerce_to_hyp_list env v = +let coerce_to_hyp_list env sigma v = let v = Value.to_list v in match v with | Some l -> - let map n = coerce_to_hyp env n in + let map n = coerce_to_hyp env sigma n in List.map map l | None -> raise (CannotCoerceTo "a variable list") (* Interprets a qualified name *) -let coerce_to_reference env v = +let coerce_to_reference env sigma v = let v = Value.normalize v in match Value.to_constr v with | Some c -> begin - try Globnames.global_of_constr c + try fst (Termops.global_of_constr sigma c) with Not_found -> raise (CannotCoerceTo "a reference") end | None -> raise (CannotCoerceTo "a reference") (* Quantified named or numbered hypothesis or hypothesis in context *) (* (as in Inversion) *) -let coerce_to_quantified_hypothesis v = +let coerce_to_quantified_hypothesis sigma v = let v = Value.normalize v in if has_type v (topwit wit_intro_pattern) then let v = out_gen (topwit wit_intro_pattern) v in @@ -322,17 +323,17 @@ let coerce_to_quantified_hypothesis v = else if has_type v (topwit wit_int) then AnonHyp (out_gen (topwit wit_int) v) else match Value.to_constr v with - | Some c when isVar c -> NamedHyp (destVar c) + | Some c when isVar sigma c -> NamedHyp (destVar sigma c) | _ -> raise (CannotCoerceTo "a quantified hypothesis") (* Quantified named or numbered hypothesis or hypothesis in context *) (* (as in Inversion) *) -let coerce_to_decl_or_quant_hyp env v = +let coerce_to_decl_or_quant_hyp env sigma v = let v = Value.normalize v in if has_type v (topwit wit_int) then AnonHyp (out_gen (topwit wit_int) v) else - try coerce_to_quantified_hypothesis v + try coerce_to_quantified_hypothesis sigma v with CannotCoerceTo _ -> raise (CannotCoerceTo "a declared or quantified hypothesis") diff --git a/ltac/taccoerce.mli b/ltac/taccoerce.mli index 3049aff7e..9c4ac5265 100644 --- a/ltac/taccoerce.mli +++ b/ltac/taccoerce.mli @@ -9,6 +9,7 @@ open Util open Names open Term +open EConstr open Misctypes open Pattern open Genarg @@ -48,16 +49,16 @@ end (** {5 Coercion functions} *) -val coerce_to_constr_context : Value.t -> EConstr.constr +val coerce_to_constr_context : Value.t -> constr -val coerce_var_to_ident : bool -> Environ.env -> Value.t -> Id.t +val coerce_var_to_ident : bool -> Environ.env -> Evd.evar_map -> Value.t -> Id.t -val coerce_to_ident_not_fresh : Evd.evar_map -> Environ.env -> Value.t -> Id.t +val coerce_to_ident_not_fresh : Environ.env -> Evd.evar_map -> Value.t -> Id.t -val coerce_to_intro_pattern : Environ.env -> Value.t -> Tacexpr.delayed_open_constr intro_pattern_expr +val coerce_to_intro_pattern : Environ.env -> Evd.evar_map -> Value.t -> Tacexpr.delayed_open_constr intro_pattern_expr val coerce_to_intro_pattern_naming : - Environ.env -> Value.t -> intro_pattern_naming_expr + Environ.env -> Evd.evar_map -> Value.t -> intro_pattern_naming_expr val coerce_to_hint_base : Value.t -> string @@ -70,22 +71,22 @@ val coerce_to_uconstr : Environ.env -> Value.t -> Glob_term.closed_glob_constr val coerce_to_closed_constr : Environ.env -> Value.t -> constr val coerce_to_evaluable_ref : - Environ.env -> Value.t -> evaluable_global_reference + Environ.env -> Evd.evar_map -> Value.t -> evaluable_global_reference val coerce_to_constr_list : Environ.env -> Value.t -> constr list val coerce_to_intro_pattern_list : - Loc.t -> Environ.env -> Value.t -> Tacexpr.intro_patterns + Loc.t -> Environ.env -> Evd.evar_map -> Value.t -> Tacexpr.intro_patterns -val coerce_to_hyp : Environ.env -> Value.t -> Id.t +val coerce_to_hyp : Environ.env -> Evd.evar_map -> Value.t -> Id.t -val coerce_to_hyp_list : Environ.env -> Value.t -> Id.t list +val coerce_to_hyp_list : Environ.env -> Evd.evar_map -> Value.t -> Id.t list -val coerce_to_reference : Environ.env -> Value.t -> Globnames.global_reference +val coerce_to_reference : Environ.env -> Evd.evar_map -> Value.t -> Globnames.global_reference -val coerce_to_quantified_hypothesis : Value.t -> quantified_hypothesis +val coerce_to_quantified_hypothesis : Evd.evar_map -> Value.t -> quantified_hypothesis -val coerce_to_decl_or_quant_hyp : Environ.env -> Value.t -> quantified_hypothesis +val coerce_to_decl_or_quant_hyp : Environ.env -> Evd.evar_map -> Value.t -> quantified_hypothesis val coerce_to_int_or_var_list : Value.t -> int or_var list diff --git a/ltac/tacinterp.ml b/ltac/tacinterp.ml index b4d2b1e50..1fe6dbdd0 100644 --- a/ltac/tacinterp.ml +++ b/ltac/tacinterp.ml @@ -239,12 +239,12 @@ let pr_value env v = else if has_type v (topwit wit_constr_context) then let c = out_gen (topwit wit_constr_context) v in match env with - | Some (env,sigma) -> pr_lconstr_env env sigma (EConstr.Unsafe.to_constr c) + | Some (env,sigma) -> pr_leconstr_env env sigma c | _ -> str "a term" else if has_type v (topwit wit_constr) then let c = out_gen (topwit wit_constr) v in match env with - | Some (env,sigma) -> pr_lconstr_env env sigma c + | Some (env,sigma) -> pr_leconstr_env env sigma c | _ -> str "a term" else if has_type v (topwit wit_constr_under_binders) then let c = out_gen (topwit wit_constr_under_binders) v in @@ -282,7 +282,7 @@ let pr_inspect env expr result = (* Transforms an id into a constr if possible, or fails with Not_found *) let constr_of_id env id = - Term.mkVar (let _ = Environ.lookup_named id env in id) + EConstr.mkVar (let _ = Environ.lookup_named id env in id) (** Generic arguments : table of interpretation functions *) @@ -334,7 +334,7 @@ let (+++) lfun1 lfun2 = Id.Map.fold Id.Map.add lfun1 lfun2 let extend_values_with_bindings (ln,lm) lfun = let of_cub c = match c with - | [], c -> Value.of_constr (EConstr.Unsafe.to_constr c) + | [], c -> Value.of_constr c | _ -> in_gen (topwit wit_constr_under_binders) c in (* For compatibility, bound variables are visible only if no other @@ -385,7 +385,7 @@ let interp_ltac_var coerce ist env locid = with Not_found -> anomaly (str "Detected '" ++ Id.print (snd locid) ++ str "' as ltac var at interning time") let interp_ident ist env sigma id = - try try_interp_ltac_var (coerce_var_to_ident false env) ist (Some (env,sigma)) (dloc,id) + try try_interp_ltac_var (coerce_var_to_ident false env sigma) ist (Some (env,sigma)) (dloc,id) with Not_found -> id (* Interprets an optional identifier, bound or fresh *) @@ -394,11 +394,11 @@ let interp_name ist env sigma = function | Name id -> Name (interp_ident ist env sigma id) let interp_intro_pattern_var loc ist env sigma id = - try try_interp_ltac_var (coerce_to_intro_pattern env) ist (Some (env,sigma)) (loc,id) + try try_interp_ltac_var (coerce_to_intro_pattern env sigma) ist (Some (env,sigma)) (loc,id) with Not_found -> IntroNaming (IntroIdentifier id) let interp_intro_pattern_naming_var loc ist env sigma id = - try try_interp_ltac_var (coerce_to_intro_pattern_naming env) ist (Some (env,sigma)) (loc,id) + try try_interp_ltac_var (coerce_to_intro_pattern_naming env sigma) ist (Some (env,sigma)) (loc,id) with Not_found -> IntroIdentifier id let interp_int ist locid = @@ -423,14 +423,14 @@ let interp_int_or_var_list ist l = (* Interprets a bound variable (especially an existing hypothesis) *) let interp_hyp ist env sigma (loc,id as locid) = (* Look first in lfun for a value coercible to a variable *) - try try_interp_ltac_var (coerce_to_hyp env) ist (Some (env,sigma)) locid + try try_interp_ltac_var (coerce_to_hyp env sigma) ist (Some (env,sigma)) locid with Not_found -> (* Then look if bound in the proof context at calling time *) if is_variable env id then id else Loc.raise ~loc (Logic.RefinerError (Logic.NoSuchHyp id)) let interp_hyp_list_as_list ist env sigma (loc,id as x) = - try coerce_to_hyp_list env (Id.Map.find id ist.lfun) + try coerce_to_hyp_list env sigma (Id.Map.find id ist.lfun) with Not_found | CannotCoerceTo _ -> [interp_hyp ist env sigma x] let interp_hyp_list ist env sigma l = @@ -445,7 +445,7 @@ let interp_move_location ist env sigma = function let interp_reference ist env sigma = function | ArgArg (_,r) -> r | ArgVar (loc, id) -> - try try_interp_ltac_var (coerce_to_reference env) ist (Some (env,sigma)) (loc, id) + try try_interp_ltac_var (coerce_to_reference env sigma) ist (Some (env,sigma)) (loc, id) with Not_found -> try VarRef (get_id (Environ.lookup_named id env)) @@ -469,7 +469,7 @@ let interp_evaluable ist env sigma = function end | ArgArg (r,None) -> r | ArgVar (loc, id) -> - try try_interp_ltac_var (coerce_to_evaluable_ref env) ist (Some (env,sigma)) (loc, id) + try try_interp_ltac_var (coerce_to_evaluable_ref env sigma) ist (Some (env,sigma)) (loc, id) with Not_found -> try try_interp_evaluable env (loc, id) with Not_found -> error_global_not_found ~loc (qualid_of_ident id) @@ -540,7 +540,7 @@ let default_fresh_id = Id.of_string "H" let interp_fresh_id ist env sigma l = let extract_ident ist env sigma id = - try try_interp_ltac_var (coerce_to_ident_not_fresh sigma env) + try try_interp_ltac_var (coerce_to_ident_not_fresh env sigma) ist (Some (env,sigma)) (dloc,id) with Not_found -> id in let ids = List.map_filter (function ArgVar (_, id) -> Some id | _ -> None) l in @@ -561,24 +561,24 @@ let interp_fresh_id ist env sigma l = Tactics.fresh_id_in_env avoid id env (* Extract the uconstr list from lfun *) -let extract_ltac_constr_context ist env = +let extract_ltac_constr_context ist env sigma = let open Glob_term in - let add_uconstr id env v map = + let add_uconstr id v map = try Id.Map.add id (coerce_to_uconstr env v) map with CannotCoerceTo _ -> map in - let add_constr id env v map = + let add_constr id v map = try Id.Map.add id (coerce_to_constr env v) map with CannotCoerceTo _ -> map in - let add_ident id env v map = - try Id.Map.add id (coerce_var_to_ident false env v) map + let add_ident id v map = + try Id.Map.add id (coerce_var_to_ident false env sigma v) map with CannotCoerceTo _ -> map in let fold id v {idents;typed;untyped} = - let idents = add_ident id env v idents in - let typed = add_constr id env v typed in - let untyped = add_uconstr id env v untyped in + let idents = add_ident id v idents in + let typed = add_constr id v typed in + let untyped = add_uconstr id v untyped in { idents ; typed ; untyped } in let empty = { idents = Id.Map.empty ;typed = Id.Map.empty ; untyped = Id.Map.empty } in @@ -586,11 +586,11 @@ let extract_ltac_constr_context ist env = (** Significantly simpler than [interp_constr], to interpret an untyped constr, it suffices to adjoin a closure environment. *) -let interp_uconstr ist env = function +let interp_uconstr ist env sigma = function | (term,None) -> - { closure = extract_ltac_constr_context ist env ; term } + { closure = extract_ltac_constr_context ist env sigma; term } | (_,Some ce) -> - let ( {typed ; untyped } as closure) = extract_ltac_constr_context ist env in + let ( {typed ; untyped } as closure) = extract_ltac_constr_context ist env sigma in let ltacvars = { Constrintern.ltac_vars = Id.(Set.union (Map.domain typed) (Map.domain untyped)); ltac_bound = Id.Map.domain ist.lfun; @@ -598,7 +598,7 @@ let interp_uconstr ist env = function { closure ; term = intern_gen WithoutTypeConstraint ~ltacvars env ce } let interp_gen kind ist allow_patvar flags env sigma (c,ce) = - let constrvars = extract_ltac_constr_context ist env in + let constrvars = extract_ltac_constr_context ist env sigma in let vars = { Pretyping.ltac_constrs = constrvars.typed; Pretyping.ltac_uconstrs = constrvars.untyped; @@ -636,10 +636,11 @@ let interp_gen kind ist allow_patvar flags env sigma (c,ce) = let (evd,c) = catch_error trace (understand_ltac flags env sigma vars kind) c in + let c = EConstr.of_constr c in (* spiwack: to avoid unnecessary modifications of tacinterp, as this function already use effect, I call [run] hoping it doesn't mess up with any assumption. *) - Proofview.NonLogical.run (db_constr (curr_debug ist) env c); + Proofview.NonLogical.run (db_constr (curr_debug ist) env evd c); (evd,c) let constr_flags = { @@ -691,7 +692,7 @@ let interp_pure_open_constr ist = let interp_typed_pattern ist env sigma (_,c,_) = let sigma, c = interp_gen WithoutTypeConstraint ist true pure_open_constr_flags env sigma c in - pattern_of_constr env sigma (EConstr.of_constr c) + pattern_of_constr env sigma c (* Interprets a constr expression *) let interp_constr_in_compound_list inj_fun dest_fun interp_fun ist env sigma l = @@ -733,10 +734,10 @@ let interp_closed_typed_pattern_with_occurrences ist env sigma (occs, a) = prioritary to an evaluable reference and otherwise to a constr (it is an encoding to satisfy the "union" type given to Simpl) *) let coerce_eval_ref_or_constr x = - try Inl (coerce_to_evaluable_ref env x) + try Inl (coerce_to_evaluable_ref env sigma x) with CannotCoerceTo _ -> let c = coerce_to_closed_constr env x in - Inr (pattern_of_constr env sigma (EConstr.of_constr c)) in + Inr (pattern_of_constr env sigma c) in (try try_interp_ltac_var coerce_eval_ref_or_constr ist (Some (env,sigma)) (loc,id) with Not_found -> error_global_not_found ~loc (qualid_of_ident id)) @@ -746,12 +747,11 @@ let interp_closed_typed_pattern_with_occurrences ist env sigma (occs, a) = let interp_constr_with_occurrences_and_name_as_list = interp_constr_in_compound_list - (fun c -> ((AllOccurrences,EConstr.of_constr c),Anonymous)) + (fun c -> ((AllOccurrences,c),Anonymous)) (function ((occs,c),Anonymous) when occs == AllOccurrences -> c | _ -> raise Not_found) (fun ist env sigma (occ_c,na) -> let (sigma,c_interp) = interp_constr_with_occurrences ist env sigma occ_c in - let c_interp = (fst c_interp, EConstr.of_constr (snd c_interp)) in sigma, (c_interp, interp_name ist env sigma na)) @@ -784,8 +784,7 @@ let interp_may_eval f ist env sigma = function let (sigma,c_interp) = f ist env sigma c in let (redfun, _) = Redexpr.reduction_of_red_expr env redexp in let sigma = Sigma.Unsafe.of_evar_map sigma in - let Sigma (c, sigma, _) = redfun.Reductionops.e_redfun env sigma (EConstr.of_constr c_interp) in - let c = EConstr.Unsafe.to_constr c in + let Sigma (c, sigma, _) = redfun.Reductionops.e_redfun env sigma c_interp in (Sigma.to_evar_map sigma, c) | ConstrContext ((loc,s),c) -> (try @@ -793,8 +792,10 @@ let interp_may_eval f ist env sigma = function let ctxt = coerce_to_constr_context (Id.Map.find s ist.lfun) in let ctxt = EConstr.Unsafe.to_constr ctxt in let evdref = ref sigma in + let ic = EConstr.Unsafe.to_constr ic in let c = subst_meta [Constr_matching.special_meta,ic] ctxt in let c = Typing.e_solve_evars env evdref (EConstr.of_constr c) in + let c = EConstr.of_constr c in !evdref , c with | Not_found -> @@ -802,8 +803,8 @@ let interp_may_eval f ist env sigma = function (str "Unbound context identifier" ++ pr_id s ++ str".")) | ConstrTypeOf c -> let (sigma,c_interp) = f ist env sigma c in - let (sigma, t) = Typing.type_of ~refresh:true env sigma (EConstr.of_constr c_interp) in - (sigma, EConstr.Unsafe.to_constr t) + let (sigma, t) = Typing.type_of ~refresh:true env sigma c_interp in + (sigma, t) | ConstrTerm c -> try f ist env sigma c @@ -833,7 +834,7 @@ let interp_constr_may_eval ist env sigma c = (* spiwack: to avoid unnecessary modifications of tacinterp, as this function already use effect, I call [run] hoping it doesn't mess up with any assumption. *) - Proofview.NonLogical.run (db_constr (curr_debug ist) env csr); + Proofview.NonLogical.run (db_constr (curr_debug ist) env sigma csr); sigma , csr end @@ -845,7 +846,7 @@ let rec message_of_value v = Ftactic.return (str "") else if has_type v (topwit wit_constr) then let v = out_gen (topwit wit_constr) v in - Ftactic.nf_enter {enter = begin fun gl -> Ftactic.return (pr_constr_env (pf_env gl) (project gl) v) end } + Ftactic.nf_enter {enter = begin fun gl -> Ftactic.return (pr_econstr_env (pf_env gl) (project gl) v) end } else if has_type v (topwit wit_constr_under_binders) then let c = out_gen (topwit wit_constr_under_binders) v in Ftactic.nf_enter { enter = begin fun gl -> @@ -922,7 +923,6 @@ and interp_intro_pattern_action ist env sigma = function let c = { delayed = fun env sigma -> let sigma = Sigma.to_evar_map sigma in let (sigma, c) = interp_open_constr ist env sigma c in - let c = EConstr.of_constr c in Sigma.Unsafe.of_pair (c, sigma) } in let sigma,ipat = interp_intro_pattern ist env sigma ipat in @@ -939,7 +939,7 @@ and interp_or_and_intro_pattern ist env sigma = function and interp_intro_pattern_list_as_list ist env sigma = function | [loc,IntroNaming (IntroIdentifier id)] as l -> - (try sigma, coerce_to_intro_pattern_list loc env (Id.Map.find id ist.lfun) + (try sigma, coerce_to_intro_pattern_list loc env sigma (Id.Map.find id ist.lfun) with Not_found | CannotCoerceTo _ -> List.fold_map (interp_intro_pattern ist env) sigma l) | l -> List.fold_map (interp_intro_pattern ist env) sigma l @@ -951,7 +951,7 @@ let interp_intro_pattern_naming_option ist env sigma = function let interp_or_and_intro_pattern_option ist env sigma = function | None -> sigma, None | Some (ArgVar (loc,id)) -> - (match coerce_to_intro_pattern env (Id.Map.find id ist.lfun) with + (match coerce_to_intro_pattern env sigma (Id.Map.find id ist.lfun) with | IntroAction (IntroOrAndPattern l) -> sigma, Some (loc,l) | _ -> raise (CannotCoerceTo "a disjunctive/conjunctive introduction pattern")) @@ -969,31 +969,25 @@ let interp_in_hyp_as ist env sigma (id,ipat) = let sigma, ipat = interp_intro_pattern_option ist env sigma ipat in sigma,(interp_hyp ist env sigma id,ipat) -let interp_quantified_hypothesis ist = function - | AnonHyp n -> AnonHyp n - | NamedHyp id -> - try try_interp_ltac_var coerce_to_quantified_hypothesis ist None(dloc,id) - with Not_found -> NamedHyp id - -let interp_binding_name ist = function +let interp_binding_name ist sigma = function | AnonHyp n -> AnonHyp n | NamedHyp id -> (* If a name is bound, it has to be a quantified hypothesis *) (* user has to use other names for variables if these ones clash with *) (* a name intented to be used as a (non-variable) identifier *) - try try_interp_ltac_var coerce_to_quantified_hypothesis ist None(dloc,id) + try try_interp_ltac_var (coerce_to_quantified_hypothesis sigma) ist None(dloc,id) with Not_found -> NamedHyp id let interp_declared_or_quantified_hypothesis ist env sigma = function | AnonHyp n -> AnonHyp n | NamedHyp id -> try try_interp_ltac_var - (coerce_to_decl_or_quant_hyp env) ist (Some (env,sigma)) (dloc,id) + (coerce_to_decl_or_quant_hyp env sigma) ist (Some (env,sigma)) (dloc,id) with Not_found -> NamedHyp id let interp_binding ist env sigma (loc,b,c) = let sigma, c = interp_open_constr ist env sigma c in - sigma, (loc,interp_binding_name ist b,c) + sigma, (loc,interp_binding_name ist sigma b,c) let interp_bindings ist env sigma = function | NoBindings -> @@ -1008,14 +1002,11 @@ let interp_bindings ist env sigma = function let interp_constr_with_bindings ist env sigma (c,bl) = let sigma, bl = interp_bindings ist env sigma bl in let sigma, c = interp_open_constr ist env sigma c in - let c = EConstr.of_constr c in - let bl = Miscops.map_bindings EConstr.of_constr bl in sigma, (c,bl) let interp_open_constr_with_bindings ist env sigma (c,bl) = let sigma, bl = interp_bindings ist env sigma bl in let sigma, c = interp_open_constr ist env sigma c in - let (c, bl) = Miscops.map_with_bindings EConstr.of_constr (c, bl) in sigma, (c, bl) let loc_of_bindings = function @@ -1053,7 +1044,7 @@ let interp_destruction_arg ist gl arg = then keep,ElimOnIdent (loc,id') else (keep, ElimOnConstr { delayed = begin fun env sigma -> - try Sigma.here (EConstr.of_constr (constr_of_id env id'), NoBindings) sigma + try Sigma.here (constr_of_id env id', NoBindings) sigma with Not_found -> user_err ~loc ~hdr:"interp_destruction_arg" ( pr_id id ++ strbrk " binds to " ++ pr_id id' ++ strbrk " which is neither a declared nor a quantified hypothesis.") @@ -1075,7 +1066,7 @@ let interp_destruction_arg ist gl arg = keep,ElimOnAnonHyp (out_gen (topwit wit_int) v) else match Value.to_constr v with | None -> error () - | Some c -> keep,ElimOnConstr { delayed = fun env sigma -> Sigma ((EConstr.of_constr c,NoBindings), sigma, Sigma.refl) } + | Some c -> keep,ElimOnConstr { delayed = fun env sigma -> Sigma ((c,NoBindings), sigma, Sigma.refl) } with Not_found -> (* We were in non strict (interactive) mode *) if Tactics.is_quantified_hypothesis id gl then @@ -1085,7 +1076,6 @@ let interp_destruction_arg ist gl arg = let f = { delayed = fun env sigma -> let sigma = Sigma.to_evar_map sigma in let (sigma,c) = interp_open_constr ist env sigma c in - let c = EConstr.of_constr c in Sigma.Unsafe.of_pair ((c,NoBindings), sigma) } in keep,ElimOnConstr f @@ -1365,7 +1355,7 @@ and interp_tacarg ist arg : Val.t Ftactic.t = Ftactic.s_enter { s_enter = begin fun gl -> let sigma = Proofview.Goal.sigma gl in let env = Proofview.Goal.env gl in - let c = interp_uconstr ist env c in + let c = interp_uconstr ist env (Sigma.to_evar_map sigma) c in let Sigma (c, sigma, p) = (type_uconstr ist c).delayed env sigma in Sigma (Ftactic.return (Value.of_constr c), sigma, p) end } @@ -1468,7 +1458,7 @@ and interp_letin ist llc u = and interp_match_success ist { Tactic_matching.subst ; context ; terms ; lhs } = let (>>=) = Ftactic.bind in let lctxt = Id.Map.map interp_context context in - let hyp_subst = Id.Map.map (EConstr.Unsafe.to_constr %> Value.of_constr) terms in + let hyp_subst = Id.Map.map Value.of_constr terms in let lfun = extend_values_with_bindings subst (lctxt +++ hyp_subst +++ ist.lfun) in let ist = { ist with lfun } in val_interp ist lhs >>= fun v -> @@ -1522,7 +1512,6 @@ and interp_match ist lz constr lmr = Proofview.tclZERO ~info e end end >>= fun constr -> - let constr = EConstr.of_constr constr in Ftactic.enter { enter = begin fun gl -> let sigma = project gl in let env = Proofview.Goal.env gl in @@ -1597,7 +1586,7 @@ and interp_genarg_var_list ist x = end } (* Interprets tactic expressions : returns a "constr" *) -and interp_ltac_constr ist e : constr Ftactic.t = +and interp_ltac_constr ist e : EConstr.t Ftactic.t = let (>>=) = Ftactic.bind in begin Proofview.tclORELSE (val_interp ist e) @@ -1625,7 +1614,7 @@ and interp_ltac_constr ist e : constr Ftactic.t = debugging_step ist (fun () -> Pptactic.pr_glob_tactic env e ++ fnl() ++ str " has value " ++ fnl() ++ - pr_constr_env env sigma cresult) + pr_econstr_env env sigma cresult) end <*> Ftactic.return cresult with CannotCoerceTo _ -> @@ -1645,7 +1634,8 @@ and name_atomic ?env tacexpr tac : unit Proofview.tactic = | Some e -> Proofview.tclUNIT e | None -> Proofview.tclENV end >>= fun env -> - let name () = Pptactic.pr_atomic_tactic env tacexpr in + Proofview.tclEVARMAP >>= fun sigma -> + let name () = Pptactic.pr_atomic_tactic env sigma tacexpr in Proofview.Trace.name_tactic name tac (* Interprets a primitive tactic *) @@ -1712,7 +1702,7 @@ and interp_atomic ist tac : unit Proofview.tactic = let env = pf_env gl in let f sigma (id,n,c) = let (sigma,c_interp) = interp_type ist env sigma c in - sigma , (interp_ident ist env sigma id,n,EConstr.of_constr c_interp) in + sigma , (interp_ident ist env sigma id,n,c_interp) in let (sigma,l_interp) = Evd.MonadR.List.map_right (fun c sigma -> f sigma c) l (project gl) in @@ -1727,7 +1717,7 @@ and interp_atomic ist tac : unit Proofview.tactic = let env = pf_env gl in let f sigma (id,c) = let (sigma,c_interp) = interp_type ist env sigma c in - sigma , (interp_ident ist env sigma id,EConstr.of_constr c_interp) in + sigma , (interp_ident ist env sigma id,c_interp) in let (sigma,l_interp) = Evd.MonadR.List.map_right (fun c sigma -> f sigma c) l (project gl) in @@ -1742,7 +1732,6 @@ and interp_atomic ist tac : unit Proofview.tactic = let (sigma,c) = (if Option.is_empty t then interp_constr else interp_type) ist env sigma c in - let c = EConstr.of_constr c in let sigma, ipat' = interp_intro_pattern_option ist env sigma ipat in let tac = Option.map (Option.map (interp_tactic ist)) t in Tacticals.New.tclWITHHOLES false @@ -1770,7 +1759,6 @@ and interp_atomic ist tac : unit Proofview.tactic = if Locusops.is_nowhere clp then (* We try to fully-typecheck the term *) let (sigma,c_interp) = interp_constr ist env sigma c in - let c_interp = EConstr.of_constr c_interp in let let_tac b na c cl eqpat = let id = Option.default (Loc.ghost,IntroAnonymous) eqpat in let with_eq = if b then None else Some (true,id) in @@ -1789,7 +1777,6 @@ and interp_atomic ist tac : unit Proofview.tactic = Tactics.letin_pat_tac with_eq na c cl in let (sigma',c) = interp_pure_open_constr ist env sigma c in - let c = EConstr.of_constr c in name_atomic ~env (TacLetTac(na,c,clp,b,eqpat)) (Tacticals.New.tclWITHHOLES false (*in hope of a future "eset/epose"*) @@ -1849,7 +1836,7 @@ and interp_atomic ist tac : unit Proofview.tactic = in let c_interp patvars = { Sigma.run = begin fun sigma -> let lfun' = Id.Map.fold (fun id c lfun -> - Id.Map.add id (Value.of_constr (EConstr.Unsafe.to_constr c)) lfun) + Id.Map.add id (Value.of_constr c) lfun) patvars ist.lfun in let sigma = Sigma.to_evar_map sigma in @@ -1859,7 +1846,6 @@ and interp_atomic ist tac : unit Proofview.tactic = then interp_type ist (pf_env gl) sigma c else interp_constr ist (pf_env gl) sigma c in - let c = EConstr.of_constr c in Sigma.Unsafe.of_pair (c, sigma) end } in Tactics.change None c_interp (interp_clause ist (pf_env gl) (project gl) cl) @@ -1876,14 +1862,13 @@ and interp_atomic ist tac : unit Proofview.tactic = let to_catch = function Not_found -> true | e -> CErrors.is_anomaly e in let c_interp patvars = { Sigma.run = begin fun sigma -> let lfun' = Id.Map.fold (fun id c lfun -> - Id.Map.add id (Value.of_constr (EConstr.Unsafe.to_constr c)) lfun) + Id.Map.add id (Value.of_constr c) lfun) patvars ist.lfun in let ist = { ist with lfun = lfun' } in try let sigma = Sigma.to_evar_map sigma in let (sigma, c) = interp_constr ist env sigma c in - let c = EConstr.of_constr c in Sigma.Unsafe.of_pair (c, sigma) with e when to_catch e (* Hack *) -> user_err (strbrk "Failed to get enough information from the left-hand side to type the right-hand side.") @@ -1922,7 +1907,6 @@ and interp_atomic ist tac : unit Proofview.tactic = | None -> sigma , None | Some c -> let (sigma,c_interp) = interp_constr ist env sigma c in - let c_interp = EConstr.of_constr c_interp in sigma , Some c_interp in let dqhyps = interp_declared_or_quantified_hypothesis ist env sigma hyp in @@ -1949,7 +1933,6 @@ and interp_atomic ist tac : unit Proofview.tactic = let env = Proofview.Goal.env gl in let sigma = project gl in let (sigma,c_interp) = interp_constr ist env sigma c in - let c_interp = EConstr.of_constr c_interp in let dqhyps = interp_declared_or_quantified_hypothesis ist env sigma hyp in let hyps = interp_hyp_list ist env sigma idl in let tac = name_atomic ~env @@ -2059,7 +2042,6 @@ end } let interp_bindings' ist bl = Ftactic.return { delayed = fun env sigma -> let (sigma, bl) = interp_bindings ist env (Sigma.to_evar_map sigma) bl in - let bl = Miscops.map_bindings EConstr.of_constr bl in Sigma.Unsafe.of_pair (bl, sigma) } @@ -2099,7 +2081,7 @@ let () = let () = register_interp0 wit_uconstr (fun ist c -> Ftactic.nf_enter { enter = begin fun gl -> - Ftactic.return (interp_uconstr ist (Proofview.Goal.env gl) c) + Ftactic.return (interp_uconstr ist (Proofview.Goal.env gl) (Tacmach.New.project gl) c) end }) (***************************************************************************) diff --git a/ltac/tacinterp.mli b/ltac/tacinterp.mli index 2b0324e24..c657a11f1 100644 --- a/ltac/tacinterp.mli +++ b/ltac/tacinterp.mli @@ -9,6 +9,7 @@ open Names open Tactic_debug open Term +open EConstr open Tacexpr open Genarg open Redexpr diff --git a/ltac/tactic_debug.ml b/ltac/tactic_debug.ml index be1123d5c..b2601ad32 100644 --- a/ltac/tactic_debug.ml +++ b/ltac/tactic_debug.ml @@ -51,8 +51,7 @@ let db_pr_goal gl = let env = Proofview.Goal.env gl in let concl = Proofview.Goal.concl gl in let penv = print_named_context env in - let concl = EConstr.Unsafe.to_constr concl in - let pc = print_constr_env env concl in + let pc = print_constr_env env (Tacmach.New.project gl) concl in str" " ++ hv 0 (penv ++ fnl () ++ str "============================" ++ fnl () ++ str" " ++ pc) ++ fnl () @@ -224,11 +223,11 @@ let is_debug db = return (Int.equal skip 0) (* Prints a constr *) -let db_constr debug env c = +let db_constr debug env sigma c = let open Proofview.NonLogical in is_debug debug >>= fun db -> if db then - msg_tac_debug (str "Evaluated term: " ++ print_constr_env env c) + msg_tac_debug (str "Evaluated term: " ++ print_constr_env env sigma c) else return () (* Prints the pattern rule *) @@ -248,20 +247,20 @@ let hyp_bound = function | Name id -> str " (bound to " ++ pr_id id ++ str ")" (* Prints a matched hypothesis *) -let db_matched_hyp debug env (id,_,c) ido = +let db_matched_hyp debug env sigma (id,_,c) ido = let open Proofview.NonLogical in is_debug debug >>= fun db -> if db then msg_tac_debug (str "Hypothesis " ++ pr_id id ++ hyp_bound ido ++ - str " has been matched: " ++ print_constr_env env c) + str " has been matched: " ++ print_constr_env env sigma c) else return () (* Prints the matched conclusion *) -let db_matched_concl debug env c = +let db_matched_concl debug env sigma c = let open Proofview.NonLogical in is_debug debug >>= fun db -> if db then - msg_tac_debug (str "Conclusion has been matched: " ++ print_constr_env env c) + msg_tac_debug (str "Conclusion has been matched: " ++ print_constr_env env sigma c) else return () (* Prints a success message when the goal has been matched *) diff --git a/ltac/tactic_debug.mli b/ltac/tactic_debug.mli index 520fb41ef..7745d9b7b 100644 --- a/ltac/tactic_debug.mli +++ b/ltac/tactic_debug.mli @@ -11,6 +11,7 @@ open Pattern open Names open Tacexpr open Term +open EConstr open Evd (** TODO: Move those definitions somewhere sensible *) @@ -34,7 +35,7 @@ val debug_prompt : val db_initialize : unit Proofview.NonLogical.t (** Prints a constr *) -val db_constr : debug_info -> env -> constr -> unit Proofview.NonLogical.t +val db_constr : debug_info -> env -> evar_map -> constr -> unit Proofview.NonLogical.t (** Prints the pattern rule *) val db_pattern_rule : @@ -42,10 +43,10 @@ val db_pattern_rule : (** Prints a matched hypothesis *) val db_matched_hyp : - debug_info -> env -> Id.t * constr option * constr -> Name.t -> unit Proofview.NonLogical.t + debug_info -> env -> evar_map -> Id.t * constr option * constr -> Name.t -> unit Proofview.NonLogical.t (** Prints the matched conclusion *) -val db_matched_concl : debug_info -> env -> constr -> unit Proofview.NonLogical.t +val db_matched_concl : debug_info -> env -> evar_map -> constr -> unit Proofview.NonLogical.t (** Prints a success message when the goal has been matched *) val db_mc_pattern_success : debug_info -> unit Proofview.NonLogical.t diff --git a/ltac/tauto.ml b/ltac/tauto.ml index a4faf438a..707154a30 100644 --- a/ltac/tauto.ml +++ b/ltac/tauto.ml @@ -25,7 +25,7 @@ let () = Mltop.add_known_module tauto_plugin let assoc_var s ist = let v = Id.Map.find (Names.Id.of_string s) ist.lfun in match Value.to_constr v with - | Some c -> EConstr.of_constr c + | Some c -> c | None -> failwith "tauto: anomaly" (** Parametrization of tauto *) diff --git a/plugins/cc/ccalgo.ml b/plugins/cc/ccalgo.ml index 102efe55b..0a980c03b 100644 --- a/plugins/cc/ccalgo.ml +++ b/plugins/cc/ccalgo.ml @@ -497,10 +497,10 @@ let rec inst_pattern subst = function args t let pr_idx_term uf i = str "[" ++ int i ++ str ":=" ++ - Termops.print_constr (constr_of_term (term uf i)) ++ str "]" + Termops.print_constr (EConstr.of_constr (constr_of_term (term uf i))) ++ str "]" let pr_term t = str "[" ++ - Termops.print_constr (constr_of_term t) ++ str "]" + Termops.print_constr (EConstr.of_constr (constr_of_term t)) ++ str "]" let rec add_term state t= let uf=state.uf in @@ -615,7 +615,7 @@ let add_inst state (inst,int_subst) = begin debug (fun () -> (str "Adding new equality, depth="++ int state.rew_depth) ++ fnl () ++ - (str " [" ++ Termops.print_constr prf ++ str " : " ++ + (str " [" ++ Termops.print_constr (EConstr.of_constr prf) ++ str " : " ++ pr_term s ++ str " == " ++ pr_term t ++ str "]")); add_equality state prf s t end @@ -623,7 +623,7 @@ let add_inst state (inst,int_subst) = begin debug (fun () -> (str "Adding new disequality, depth="++ int state.rew_depth) ++ fnl () ++ - (str " [" ++ Termops.print_constr prf ++ str " : " ++ + (str " [" ++ Termops.print_constr (EConstr.of_constr prf) ++ str " : " ++ pr_term s ++ str " <> " ++ pr_term t ++ str "]")); add_disequality state (Hyp prf) s t end diff --git a/plugins/cc/cctac.ml b/plugins/cc/cctac.ml index a4ed4798a..62892973d 100644 --- a/plugins/cc/cctac.ml +++ b/plugins/cc/cctac.ml @@ -414,6 +414,7 @@ let build_term_to_complete uf meta pac = let cc_tactic depth additionnal_terms = Proofview.Goal.nf_enter { enter = begin fun gl -> + let sigma = Tacmach.New.project gl in Coqlib.check_required_library Coqlib.logic_module_name; let _ = debug (fun () -> Pp.str "Reading subgoal ...") in let state = Tacmach.New.of_old (fun gls -> make_prb gls depth additionnal_terms) gl in @@ -448,7 +449,7 @@ let cc_tactic depth additionnal_terms = str "\"congruence with (" ++ prlist_with_sep (fun () -> str ")" ++ spc () ++ str "(") - (EConstr.Unsafe.to_constr %> Termops.print_constr_env env) + (Termops.print_constr_env env sigma) terms_to_complete ++ str ")\"," end ++ diff --git a/plugins/cc/g_congruence.ml4 b/plugins/cc/g_congruence.ml4 index b787e824f..6f6811334 100644 --- a/plugins/cc/g_congruence.ml4 +++ b/plugins/cc/g_congruence.ml4 @@ -18,9 +18,9 @@ DECLARE PLUGIN "cc_plugin" TACTIC EXTEND cc [ "congruence" ] -> [ congruence_tac 1000 [] ] |[ "congruence" integer(n) ] -> [ congruence_tac n [] ] - |[ "congruence" "with" ne_constr_list(l) ] -> [ congruence_tac 1000 (List.map EConstr.of_constr l) ] + |[ "congruence" "with" ne_constr_list(l) ] -> [ congruence_tac 1000 l ] |[ "congruence" integer(n) "with" ne_constr_list(l) ] -> - [ congruence_tac n (List.map EConstr.of_constr l) ] + [ congruence_tac n l ] END TACTIC EXTEND f_equal diff --git a/plugins/decl_mode/ppdecl_proof.ml b/plugins/decl_mode/ppdecl_proof.ml index 59a0bb5a2..93b98263e 100644 --- a/plugins/decl_mode/ppdecl_proof.ml +++ b/plugins/decl_mode/ppdecl_proof.ml @@ -206,6 +206,8 @@ let pr_glob_proof_instr pconstr1 pconstr2 ptac (instr : glob_proof_instr) = instr let pr_proof_instr pconstr1 pconstr2 ptac (instr : proof_instr) = + let pconstr1 c = pconstr1 (EConstr.of_constr c) in + let pconstr2 c = pconstr2 (EConstr.of_constr c) in pr_gen_proof_instr (fun st -> pr_statement pconstr1 st) pconstr2 diff --git a/plugins/funind/g_indfun.ml4 b/plugins/funind/g_indfun.ml4 index a6f971703..560242bf2 100644 --- a/plugins/funind/g_indfun.ml4 +++ b/plugins/funind/g_indfun.ml4 @@ -58,7 +58,7 @@ let pr_fun_ind_using_typed prc prlc _ opt_c = | None -> mt () | Some b -> let (b, _) = Tactics.run_delayed (Global.env ()) Evd.empty b in - spc () ++ hov 2 (str "using" ++ spc () ++ pr_with_bindings_typed (EConstr.Unsafe.to_constr %> prc) (EConstr.Unsafe.to_constr %> prlc) b) + spc () ++ hov 2 (str "using" ++ spc () ++ pr_with_bindings_typed prc prlc b) ARGUMENT EXTEND fun_ind_using @@ -108,8 +108,9 @@ TACTIC EXTEND newfunind let c = match cl with | [] -> assert false | [c] -> c - | c::cl -> applist(c,cl) + | c::cl -> EConstr.applist(c,cl) in + let c = EConstr.Unsafe.to_constr c in Extratactics.onSomeWithHoles (fun x -> functional_induction true c x pat) princl ] END (***** debug only ***) @@ -119,8 +120,9 @@ TACTIC EXTEND snewfunind let c = match cl with | [] -> assert false | [c] -> c - | c::cl -> applist(c,cl) + | c::cl -> EConstr.applist(c,cl) in + let c = EConstr.Unsafe.to_constr c in Extratactics.onSomeWithHoles (fun x -> functional_induction false c x pat) princl ] END diff --git a/plugins/nsatz/g_nsatz.ml4 b/plugins/nsatz/g_nsatz.ml4 index 5f906a8da..a6749c347 100644 --- a/plugins/nsatz/g_nsatz.ml4 +++ b/plugins/nsatz/g_nsatz.ml4 @@ -13,5 +13,5 @@ DECLARE PLUGIN "nsatz_plugin" DECLARE PLUGIN "nsatz_plugin" TACTIC EXTEND nsatz_compute -| [ "nsatz_compute" constr(lt) ] -> [ Nsatz.nsatz_compute lt ] +| [ "nsatz_compute" constr(lt) ] -> [ Nsatz.nsatz_compute (EConstr.Unsafe.to_constr lt) ] END diff --git a/plugins/quote/g_quote.ml4 b/plugins/quote/g_quote.ml4 index 79c429615..40c1028e5 100644 --- a/plugins/quote/g_quote.ml4 +++ b/plugins/quote/g_quote.ml4 @@ -23,7 +23,6 @@ let cont = Id.of_string "cont" let x = Id.of_string "x" let make_cont (k : Val.t) (c : EConstr.t) = - let c = EConstr.Unsafe.to_constr c in let c = Tacinterp.Value.of_constr c in let tac = TacCall (loc, ArgVar (loc, cont), [Reference (ArgVar (loc, x))]) in let ist = { lfun = Id.Map.add cont k (Id.Map.singleton x c); extra = TacStore.empty; } in @@ -33,8 +32,8 @@ TACTIC EXTEND quote [ "quote" ident(f) ] -> [ quote f [] ] | [ "quote" ident(f) "[" ne_ident_list(lc) "]"] -> [ quote f lc ] | [ "quote" ident(f) "in" constr(c) "using" tactic(k) ] -> - [ gen_quote (make_cont k) (EConstr.of_constr c) f [] ] + [ gen_quote (make_cont k) c f [] ] | [ "quote" ident(f) "[" ne_ident_list(lc) "]" "in" constr(c) "using" tactic(k) ] -> - [ gen_quote (make_cont k) (EConstr.of_constr c) f lc ] + [ gen_quote (make_cont k) c f lc ] END diff --git a/plugins/setoid_ring/g_newring.ml4 b/plugins/setoid_ring/g_newring.ml4 index 13cf8330b..b1882ae8a 100644 --- a/plugins/setoid_ring/g_newring.ml4 +++ b/plugins/setoid_ring/g_newring.ml4 @@ -30,7 +30,7 @@ END TACTIC EXTEND closed_term [ "closed_term" constr(t) "[" ne_reference_list(l) "]" ] -> - [ closed_term (EConstr.of_constr t) l ] + [ closed_term t l ] END open Pptactic @@ -90,11 +90,7 @@ END TACTIC EXTEND ring_lookup | [ "ring_lookup" tactic0(f) "[" constr_list(lH) "]" ne_constr_list(lrt) ] -> - [ - let lH = List.map EConstr.of_constr lH in - let lrt = List.map EConstr.of_constr lrt in - let (t,lr) = List.sep_last lrt in ring_lookup f lH lr t - ] + [ let (t,lr) = List.sep_last lrt in ring_lookup f lH lr t ] END let pr_field_mod = function @@ -129,9 +125,5 @@ END TACTIC EXTEND field_lookup | [ "field_lookup" tactic(f) "[" constr_list(lH) "]" ne_constr_list(lt) ] -> - [ - let lH = List.map EConstr.of_constr lH in - let lt = List.map EConstr.of_constr lt in - let (t,l) = List.sep_last lt in field_lookup f lH l t - ] + [ let (t,l) = List.sep_last lt in field_lookup f lH l t ] END diff --git a/plugins/setoid_ring/newring.ml b/plugins/setoid_ring/newring.ml index c0eeff8d7..ce2c558ae 100644 --- a/plugins/setoid_ring/newring.ml +++ b/plugins/setoid_ring/newring.ml @@ -183,7 +183,7 @@ let dummy_goal env sigma = {Evd.it = gl; Evd.sigma = sigma} let constr_of v = match Value.to_constr v with - | Some c -> c + | Some c -> EConstr.Unsafe.to_constr c | None -> failwith "Ring.exec_tactic: anomaly" let tactic_res = ref [||] @@ -203,7 +203,6 @@ let get_res = let exec_tactic env evd n f args = let fold arg (i, vars, lfun) = - let arg = EConstr.Unsafe.to_constr arg in let id = Id.of_string ("x" ^ string_of_int i) in let x = Reference (ArgVar (Loc.ghost, id)) in (succ i, x :: vars, Id.Map.add id (Value.of_constr arg) lfun) @@ -730,7 +729,7 @@ let make_term_list env evd carrier rl = (plapp evd coq_nil [|carrier|]) in Typing.e_solve_evars env evd l -let carg = Tacinterp.Value.of_constr +let carg c = Tacinterp.Value.of_constr (EConstr.of_constr c) let tacarg expr = Tacinterp.Value.of_closure (Tacinterp.default_ist ()) expr diff --git a/plugins/ssrmatching/ssrmatching.ml4 b/plugins/ssrmatching/ssrmatching.ml4 index 9798fa11c..9dcc6c4cc 100644 --- a/plugins/ssrmatching/ssrmatching.ml4 +++ b/plugins/ssrmatching/ssrmatching.ml4 @@ -1030,7 +1030,7 @@ let interp_constr = interp_wit wit_constr let interp_open_constr ist gl gc = interp_wit wit_open_constr ist gl gc let pf_intern_term ist gl (_, c) = glob_constr ist (pf_env gl) c -let interp_term ist gl (_, c) = (interp_open_constr ist gl c) +let interp_term ist gl (_, c) = on_snd EConstr.Unsafe.to_constr (interp_open_constr ist gl c) let pr_ssrterm _ _ _ = pr_term let input_ssrtermkind strm = match Compat.get_tok (stream_nth 0 strm) with | Tok.KEYWORD "(" -> '(' diff --git a/pretyping/evardefine.ml b/pretyping/evardefine.ml index d4b46c046..875e4a118 100644 --- a/pretyping/evardefine.ml +++ b/pretyping/evardefine.ml @@ -208,6 +208,6 @@ let split_tycon loc env evd tycon = let valcon_of_tycon x = x let lift_tycon n = Option.map (lift n) -let pr_tycon env = function +let pr_tycon env sigma = function None -> str "None" - | Some t -> Termops.print_constr_env env (EConstr.Unsafe.to_constr t) + | Some t -> Termops.print_constr_env env sigma t diff --git a/pretyping/evardefine.mli b/pretyping/evardefine.mli index 9c03a6e3f..2f7ac4efb 100644 --- a/pretyping/evardefine.mli +++ b/pretyping/evardefine.mli @@ -43,5 +43,5 @@ val define_evar_as_sort : env -> evar_map -> existential -> evar_map * sorts (** {6 debug pretty-printer:} *) -val pr_tycon : env -> type_constraint -> Pp.std_ppcmds +val pr_tycon : env -> evar_map -> type_constraint -> Pp.std_ppcmds diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml index a0d8faab4..09b99983c 100644 --- a/pretyping/pretyping.ml +++ b/pretyping/pretyping.ml @@ -429,7 +429,7 @@ let protected_get_type_of env sigma c = try Retyping.get_type_of ~lax:true env.ExtraEnv.env sigma c with Retyping.RetypeError _ -> user_err - (str "Cannot reinterpret " ++ quote (print_constr (EConstr.Unsafe.to_constr c)) ++ + (str "Cannot reinterpret " ++ quote (print_constr c) ++ str " in the current environment.") let pretype_id pretype k0 loc env evdref lvar id = @@ -1225,6 +1225,7 @@ let type_uconstr ?(flags = constr_flags) } in let sigma = Sigma.to_evar_map sigma in let (sigma, c) = understand_ltac flags env sigma vars expected_type term in + let c = EConstr.of_constr c in Sigma.Unsafe.of_pair (c, sigma) end } diff --git a/pretyping/pretyping.mli b/pretyping/pretyping.mli index 2f3ce3afa..a1602088a 100644 --- a/pretyping/pretyping.mli +++ b/pretyping/pretyping.mli @@ -119,7 +119,7 @@ val understand_judgment_tcc : env -> evar_map ref -> val type_uconstr : ?flags:inference_flags -> ?expected_type:typing_constraint -> - Geninterp.interp_sign -> Glob_term.closed_glob_constr -> constr Tactypes.delayed_open + Geninterp.interp_sign -> Glob_term.closed_glob_constr -> EConstr.constr Tactypes.delayed_open (** Trying to solve remaining evars and remaining conversion problems possibly using type classes, heuristics, external tactic solver diff --git a/pretyping/recordops.ml b/pretyping/recordops.ml index 8362a2a26..bc9e3a1f4 100644 --- a/pretyping/recordops.ml +++ b/pretyping/recordops.ml @@ -191,7 +191,7 @@ let warn_projection_no_head_constant = CWarnings.create ~name:"projection-no-head-constant" ~category:"typechecker" (fun (t,con_pp,proji_sp_pp) -> strbrk "Projection value has no head constant: " - ++ Termops.print_constr t ++ strbrk " in canonical instance " + ++ Termops.print_constr (EConstr.of_constr t) ++ strbrk " in canonical instance " ++ con_pp ++ str " of " ++ proji_sp_pp ++ strbrk ", ignoring it.") (* Intended to always succeed *) @@ -256,8 +256,8 @@ let add_canonical_structure warn o = in match ocs with | None -> object_table := Refmap.add proj ((pat,s)::l) !object_table; | Some (c, cs) -> - let old_can_s = (Termops.print_constr cs.o_DEF) - and new_can_s = (Termops.print_constr s.o_DEF) in + let old_can_s = (Termops.print_constr (EConstr.of_constr cs.o_DEF)) + and new_can_s = (Termops.print_constr (EConstr.of_constr s.o_DEF)) in let prj = (Nametab.pr_global_env Id.Set.empty proj) and hd_val = (pr_cs_pattern cs_pat) in if warn then warn_redundant_canonical_projection (hd_val,prj,new_can_s,old_can_s)) diff --git a/pretyping/reductionops.ml b/pretyping/reductionops.ml index 90c5b241b..bc5c629f4 100644 --- a/pretyping/reductionops.ml +++ b/pretyping/reductionops.ml @@ -210,7 +210,7 @@ module Cst_stack = struct let pr l = let open Pp in - let p_c c = Termops.print_constr (EConstr.Unsafe.to_constr c) in + let p_c c = Termops.print_constr c in prlist_with_sep pr_semicolon (fun (c,params,args) -> hov 1 (str"(" ++ p_c c ++ str ")" ++ spc () ++ pr_sequence p_c params ++ spc () ++ str "(args:" ++ @@ -606,7 +606,7 @@ type local_state_reduction_function = evar_map -> state -> state let pr_state (tm,sk) = let open Pp in - let pr c = Termops.print_constr (EConstr.Unsafe.to_constr c) in + let pr c = Termops.print_constr c in h 0 (pr tm ++ str "|" ++ cut () ++ Stack.pr pr sk) let local_assum (na, t) = @@ -835,7 +835,7 @@ let rec whd_state_gen ?csts ~refold ~tactic_mode flags env sigma = let rec whrec cst_l (x, stack) = let () = if !debug_RAKAM then let open Pp in - let pr c = Termops.print_constr (EConstr.Unsafe.to_constr c) in + let pr c = Termops.print_constr c in Feedback.msg_notice (h 0 (str "<<" ++ pr x ++ str "|" ++ cut () ++ Cst_stack.pr cst_l ++ diff --git a/pretyping/unification.ml b/pretyping/unification.ml index 233b58e91..c6fad1a34 100644 --- a/pretyping/unification.ml +++ b/pretyping/unification.ml @@ -671,16 +671,13 @@ let eta_constructor_app env sigma f l1 term = | _ -> assert false) | _ -> assert false -let print_constr_env env c = - print_constr_env env (EConstr.Unsafe.to_constr c) - let rec unify_0_with_initial_metas (sigma,ms,es as subst : subst0) conv_at_top env cv_pb flags m n = let rec unirec_rec (curenv,nb as curenvnb) pb opt ((sigma,metasubst,evarsubst) as substn : subst0) curm curn = let cM = Evarutil.whd_head_evar sigma curm and cN = Evarutil.whd_head_evar sigma curn in let () = if !debug_unification then - Feedback.msg_debug (print_constr_env curenv cM ++ str" ~= " ++ print_constr_env curenv cN) + Feedback.msg_debug (print_constr_env curenv sigma cM ++ str" ~= " ++ print_constr_env curenv sigma cN) in match (EConstr.kind sigma cM, EConstr.kind sigma cN) with | Meta k1, Meta k2 -> diff --git a/printing/printer.ml b/printing/printer.ml index 2a30681bf..4a6c83bd7 100644 --- a/printing/printer.ml +++ b/printing/printer.ml @@ -60,7 +60,10 @@ let pr_constr_goal_style_env env = pr_constr_core true env let pr_open_lconstr_env env sigma (_,c) = pr_lconstr_env env sigma c let pr_open_constr_env env sigma (_,c) = pr_constr_env env sigma c - (* NB do not remove the eta-redexes! Global.env() has side-effects... *) +let pr_leconstr_env env sigma c = pr_lconstr_env env sigma (EConstr.to_constr sigma c) +let pr_econstr_env env sigma c = pr_constr_env env sigma (EConstr.to_constr sigma c) + +(* NB do not remove the eta-redexes! Global.env() has side-effects... *) let pr_lconstr t = let (sigma, env) = get_current_context () in pr_lconstr_env env sigma t @@ -71,15 +74,18 @@ let pr_constr t = let pr_open_lconstr (_,c) = pr_lconstr c let pr_open_constr (_,c) = pr_constr c +let pr_leconstr c = pr_lconstr (EConstr.Unsafe.to_constr c) +let pr_econstr c = pr_constr (EConstr.Unsafe.to_constr c) + let pr_constr_under_binders_env_gen pr env sigma (ids,c) = (* Warning: clashes can occur with variables of same name in env but *) (* we also need to preserve the actual names of the patterns *) (* So what to do? *) let assums = List.map (fun id -> (Name id,(* dummy *) mkProp)) ids in - pr (Termops.push_rels_assum assums env) sigma (EConstr.Unsafe.to_constr c) + pr (Termops.push_rels_assum assums env) sigma c -let pr_constr_under_binders_env = pr_constr_under_binders_env_gen pr_constr_env -let pr_lconstr_under_binders_env = pr_constr_under_binders_env_gen pr_lconstr_env +let pr_constr_under_binders_env = pr_constr_under_binders_env_gen pr_econstr_env +let pr_lconstr_under_binders_env = pr_constr_under_binders_env_gen pr_leconstr_env let pr_constr_under_binders c = let (sigma, env) = get_current_context () in @@ -147,7 +153,7 @@ let pr_constr_pattern t = let pr_sort sigma s = pr_glob_sort (extern_sort sigma s) let _ = Termops.set_print_constr - (fun env t -> pr_lconstr_expr (extern_constr ~lax:true false env Evd.empty t)) + (fun env sigma t -> pr_lconstr_expr (extern_constr ~lax:true false env sigma (EConstr.Unsafe.to_constr t))) let pr_in_comment pr x = str "(* " ++ pr x ++ str " *)" diff --git a/printing/printer.mli b/printing/printer.mli index 20032012a..7521468e2 100644 --- a/printing/printer.mli +++ b/printing/printer.mli @@ -38,6 +38,10 @@ val safe_pr_lconstr : constr -> std_ppcmds val safe_pr_constr_env : env -> evar_map -> constr -> std_ppcmds val safe_pr_constr : constr -> std_ppcmds +val pr_econstr_env : env -> evar_map -> EConstr.t -> std_ppcmds +val pr_econstr : EConstr.t -> std_ppcmds +val pr_leconstr_env : env -> evar_map -> EConstr.t -> std_ppcmds +val pr_leconstr : EConstr.t -> std_ppcmds val pr_open_constr_env : env -> evar_map -> open_constr -> std_ppcmds val pr_open_constr : open_constr -> std_ppcmds diff --git a/proofs/clenv.ml b/proofs/clenv.ml index 393e958d3..7269c61e3 100644 --- a/proofs/clenv.ml +++ b/proofs/clenv.ml @@ -575,10 +575,9 @@ let make_clenv_binding env sigma = make_clenv_binding_gen false None env sigma (* Pretty-print *) let pr_clenv clenv = - let inj = EConstr.Unsafe.to_constr in h 0 - (str"TEMPL: " ++ print_constr (inj clenv.templval.rebus) ++ - str" : " ++ print_constr (inj clenv.templtyp.rebus) ++ fnl () ++ + (str"TEMPL: " ++ print_constr clenv.templval.rebus ++ + str" : " ++ print_constr clenv.templtyp.rebus ++ fnl () ++ pr_evar_map (Some 2) clenv.evd) (****************************************************************) diff --git a/proofs/redexpr.ml b/proofs/redexpr.ml index 8878051c8..0fe5c73f1 100644 --- a/proofs/redexpr.ml +++ b/proofs/redexpr.ml @@ -11,6 +11,7 @@ open CErrors open Util open Names open Term +open EConstr open Declarations open Globnames open Genredexpr @@ -200,9 +201,6 @@ let out_arg = function let out_with_occurrences (occs,c) = (Locusops.occurrences_map (List.map out_arg) occs, c) -let out_with_occurrences' (occs,c) = - (Locusops.occurrences_map (List.map out_arg) occs, EConstr.of_constr c) - let e_red f = { e_redfun = fun env evm c -> Sigma.here (f env (Sigma.to_evar_map evm) c) evm } let head_style = false (* Turn to true to have a semantics where simpl @@ -242,8 +240,8 @@ let reduction_of_red_expr env = (e_red (strong_cbn (make_flag f)), DEFAULTcast) | Lazy f -> (e_red (clos_norm_flags (make_flag f)),DEFAULTcast) | Unfold ubinds -> (e_red (unfoldn (List.map out_with_occurrences ubinds)),DEFAULTcast) - | Fold cl -> (e_red (fold_commands (List.map EConstr.of_constr cl)),DEFAULTcast) - | Pattern lp -> (pattern_occs (List.map out_with_occurrences' lp),DEFAULTcast) + | Fold cl -> (e_red (fold_commands cl),DEFAULTcast) + | Pattern lp -> (pattern_occs (List.map out_with_occurrences lp),DEFAULTcast) | ExtraRedExpr s -> (try (e_red (String.Map.find s !reduction_tab),DEFAULTcast) with Not_found -> @@ -256,9 +254,12 @@ let reduction_of_red_expr env = in reduction_of_red_expr +let subst_mps subst c = + EConstr.of_constr (Mod_subst.subst_mps subst (EConstr.Unsafe.to_constr c)) + let subst_red_expr subs = Miscops.map_red_expr_gen - (Mod_subst.subst_mps subs) + (subst_mps subs) (Mod_subst.subst_evaluable_reference subs) (Patternops.subst_pattern subs) diff --git a/proofs/redexpr.mli b/proofs/redexpr.mli index d4c2c4a6c..45e461066 100644 --- a/proofs/redexpr.mli +++ b/proofs/redexpr.mli @@ -10,6 +10,7 @@ open Names open Term +open EConstr open Pattern open Genredexpr open Reductionops diff --git a/proofs/tacmach.ml b/proofs/tacmach.ml index 7ffb30fa8..3641ad74d 100644 --- a/proofs/tacmach.ml +++ b/proofs/tacmach.ml @@ -144,7 +144,7 @@ open Pp let db_pr_goal sigma g = let env = Goal.V82.env sigma g in let penv = print_named_context env in - let pc = print_constr_env env (EConstr.Unsafe.to_constr (Goal.V82.concl sigma g)) in + let pc = print_constr_env env sigma (Goal.V82.concl sigma g) in str" " ++ hv 0 (penv ++ fnl () ++ str "============================" ++ fnl () ++ str" " ++ pc) ++ fnl () diff --git a/tactics/auto.ml b/tactics/auto.ml index 4218be0bb..b548f8b92 100644 --- a/tactics/auto.ml +++ b/tactics/auto.ml @@ -152,7 +152,7 @@ let conclPattern concl pat tac = let open Genarg in let open Geninterp in let inj c = match val_tag (topwit Stdarg.wit_constr) with - | Val.Base tag -> Val.Dyn (tag, EConstr.Unsafe.to_constr c) + | Val.Base tag -> Val.Dyn (tag, c) | _ -> assert false in let fold id c accu = Id.Map.add id (inj c) accu in diff --git a/tactics/tactics.ml b/tactics/tactics.ml index 574f1c6f3..4e833eb55 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -51,7 +51,7 @@ open Context.Named.Declaration module RelDecl = Context.Rel.Declaration module NamedDecl = Context.Named.Declaration -let inj_with_occurrences e = (AllOccurrences,EConstr.Unsafe.to_constr e) +let inj_with_occurrences e = (AllOccurrences,e) let dloc = Loc.ghost @@ -922,7 +922,7 @@ let reduction_clause redexp cl = let reduce redexp cl = let trace () = let open Printer in - let pr = (pr_constr, pr_lconstr, pr_evaluable_reference, pr_constr_pattern) in + let pr = (pr_econstr, pr_leconstr, pr_evaluable_reference, pr_constr_pattern) in Pp.(hov 2 (Pputils.pr_red_expr pr str redexp)) in Proofview.Trace.name_tactic trace begin -- cgit v1.2.3 From 778e863b77bcafc8ed339dd02226e85e5fee2532 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 25 Nov 2016 11:36:09 +0100 Subject: Removing compatibility layers related to printing. --- ide/ide_slave.ml | 1 + ltac/rewrite.ml | 9 +++-- ltac/tacinterp.ml | 8 +++-- plugins/decl_mode/g_decl_mode.ml4 | 2 +- plugins/setoid_ring/newring.ml | 2 +- pretyping/evarconv.ml | 2 +- printing/prettyp.ml | 15 +++++---- printing/prettyp.mli | 8 ++--- printing/printer.ml | 12 ++++--- printing/printer.mli | 9 +++-- tactics/class_tactics.ml | 9 +++-- tactics/hints.ml | 8 ++--- tactics/leminv.ml | 8 ++--- tactics/tactics.ml | 2 +- toplevel/himsg.ml | 70 +++++++++++++-------------------------- toplevel/vernacentries.ml | 11 +++--- 16 files changed, 81 insertions(+), 95 deletions(-) diff --git a/ide/ide_slave.ml b/ide/ide_slave.ml index c11c78587..fcba01353 100644 --- a/ide/ide_slave.ml +++ b/ide/ide_slave.ml @@ -192,6 +192,7 @@ let process_goal sigma g = let id = Goal.uid g in let ccl = let norm_constr = Reductionops.nf_evar sigma (EConstr.Unsafe.to_constr (Goal.V82.concl sigma g)) in + let norm_constr = EConstr.of_constr norm_constr in Richpp.richpp_of_pp (pr_goal_concl_style_env env sigma norm_constr) in let process_hyp d (env,l) = diff --git a/ltac/rewrite.ml b/ltac/rewrite.ml index 85c19fac4..3c6bd4563 100644 --- a/ltac/rewrite.ml +++ b/ltac/rewrite.ml @@ -1526,7 +1526,7 @@ let cl_rewrite_clause_aux ?(abs=None) strat env avoid sigma concl is_hyp : resul | Success res -> let (_, cstrs) = res.rew_evars in let evars' = solve_constraints env res.rew_evars in - let newt = EConstr.of_constr (Evarutil.nf_evar evars' (EConstr.Unsafe.to_constr res.rew_to)) in + let newt = nf_evar evars' res.rew_to in let evars = (* Keep only original evars (potentially instantiated) and goal evars, the rest has been defined and substituted already. *) Evar.Set.fold @@ -2151,10 +2151,9 @@ let _ = Hook.set Equality.general_setoid_rewrite_clause general_s_rewrite (** [setoid_]{reflexivity,symmetry,transitivity} tactics *) -let not_declared env ty rel = - let rel = EConstr.Unsafe.to_constr rel in +let not_declared env sigma ty rel = tclFAIL 0 - (str" The relation " ++ Printer.pr_constr_env env Evd.empty rel ++ str" is not a declared " ++ + (str" The relation " ++ Printer.pr_econstr_env env sigma rel ++ str" is not a declared " ++ str ty ++ str" relation. Maybe you need to require the Coq.Classes.RelationClasses library") let setoid_proof ty fn fallback = @@ -2181,7 +2180,7 @@ let setoid_proof ty fn fallback = begin match e with | (Not_found, _) -> let rel, _, _ = decompose_app_rel env sigma concl in - not_declared env ty rel + not_declared env sigma ty rel | (e, info) -> Proofview.tclZERO ~info e end | e' -> Proofview.tclZERO ~info e' diff --git a/ltac/tacinterp.ml b/ltac/tacinterp.ml index 1fe6dbdd0..d8df07733 100644 --- a/ltac/tacinterp.ml +++ b/ltac/tacinterp.ml @@ -858,14 +858,16 @@ let rec message_of_value v = Ftactic.return (int (out_gen (topwit wit_int) v)) else if has_type v (topwit wit_intro_pattern) then let p = out_gen (topwit wit_intro_pattern) v in - let print env sigma c = pr_constr_env env sigma (EConstr.Unsafe.to_constr (fst (Tactics.run_delayed env Evd.empty c))) in + let print env sigma c = + let (c, sigma) = Tactics.run_delayed env sigma c in + pr_econstr_env env sigma c + in Ftactic.nf_enter { enter = begin fun gl -> Ftactic.return (Miscprint.pr_intro_pattern (fun c -> print (pf_env gl) (project gl) c) p) end } else if has_type v (topwit wit_constr_context) then let c = out_gen (topwit wit_constr_context) v in - let c = EConstr.Unsafe.to_constr c in - Ftactic.nf_enter { enter = begin fun gl -> Ftactic.return (pr_constr_env (pf_env gl) (project gl) c) end } + Ftactic.nf_enter { enter = begin fun gl -> Ftactic.return (pr_econstr_env (pf_env gl) (project gl) c) end } else if has_type v (topwit wit_uconstr) then let c = out_gen (topwit wit_uconstr) v in Ftactic.nf_enter { enter = begin fun gl -> diff --git a/plugins/decl_mode/g_decl_mode.ml4 b/plugins/decl_mode/g_decl_mode.ml4 index 9e2c9f597..18a35c6cf 100644 --- a/plugins/decl_mode/g_decl_mode.ml4 +++ b/plugins/decl_mode/g_decl_mode.ml4 @@ -25,7 +25,7 @@ open Ppdecl_proof let pr_goal gs = let (g,sigma) = Goal.V82.nf_evar (Tacmach.project gs) (Evd.sig_it gs) in let env = Goal.V82.env sigma g in - let concl = EConstr.Unsafe.to_constr (Goal.V82.concl sigma g) in + let concl = Goal.V82.concl sigma g in let goal = Printer.pr_context_of env sigma ++ cut () ++ str "============================" ++ cut () ++ diff --git a/plugins/setoid_ring/newring.ml b/plugins/setoid_ring/newring.ml index ce2c558ae..358ea5685 100644 --- a/plugins/setoid_ring/newring.ml +++ b/plugins/setoid_ring/newring.ml @@ -343,7 +343,7 @@ let _ = add_map "ring" (****************************************************************************) (* Ring database *) -let pr_constr c = pr_constr (EConstr.Unsafe.to_constr c) +let pr_constr c = pr_econstr c module Cmap = Map.Make(Constr) diff --git a/pretyping/evarconv.ml b/pretyping/evarconv.ml index afb0bf6d5..87267d538 100644 --- a/pretyping/evarconv.ml +++ b/pretyping/evarconv.ml @@ -385,7 +385,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty match is_unification_pattern_evar env evd ev lF tM with | None -> fallback () | Some l1' -> (* Miller-Pfenning's patterns unification *) - let t2 = EConstr.of_constr (nf_evar evd (EConstr.Unsafe.to_constr tM)) (** FIXME *) in + let t2 = tM in let t2 = solve_pattern_eqn env evd l1' t2 in solve_simple_eqn (evar_conv_x ts) env evd (position_problem on_left pbty,ev,t2) diff --git a/printing/prettyp.ml b/printing/prettyp.ml index e2c0f55f8..93970512d 100644 --- a/printing/prettyp.ml +++ b/printing/prettyp.ml @@ -42,8 +42,8 @@ type object_pr = { print_named_decl : Context.Named.Declaration.t -> std_ppcmds; print_library_entry : bool -> (object_name * Lib.node) -> std_ppcmds option; print_context : bool -> int option -> Lib.library_segment -> std_ppcmds; - print_typed_value_in_env : Environ.env -> Evd.evar_map -> Term.constr * Term.types -> Pp.std_ppcmds; - print_eval : Reductionops.reduction_function -> env -> Evd.evar_map -> Constrexpr.constr_expr -> unsafe_judgment -> std_ppcmds; + print_typed_value_in_env : Environ.env -> Evd.evar_map -> EConstr.constr * EConstr.types -> Pp.std_ppcmds; + print_eval : Reductionops.reduction_function -> env -> Evd.evar_map -> Constrexpr.constr_expr -> EConstr.unsafe_judgment -> std_ppcmds; } let gallina_print_module = print_module @@ -433,8 +433,8 @@ let print_located_qualid ref = print_located_qualid "object" [`TERM; `LTAC; `MOD (**** Gallina layer *****) let gallina_print_typed_value_in_env env sigma (trm,typ) = - (pr_lconstr_env env sigma trm ++ fnl () ++ - str " : " ++ pr_ltype_env env sigma typ) + (pr_leconstr_env env sigma trm ++ fnl () ++ + str " : " ++ pr_letype_env env sigma typ) (* To be improved; the type should be used to provide the types in the abstractions. This should be done recursively inside pr_lconstr, so that @@ -595,8 +595,7 @@ let gallina_print_context with_values = prec let gallina_print_eval red_fun env sigma _ {uj_val=trm;uj_type=typ} = - let ntrm = red_fun env sigma (EConstr.of_constr trm) in - let ntrm = EConstr.Unsafe.to_constr ntrm in + let ntrm = red_fun env sigma trm in (str " = " ++ gallina_print_typed_value_in_env env sigma (ntrm,typ)) (******************************************) @@ -643,6 +642,8 @@ let print_judgment env sigma {uj_val=trm;uj_type=typ} = let print_safe_judgment env sigma j = let trm = Safe_typing.j_val j in let typ = Safe_typing.j_type j in + let trm = EConstr.of_constr trm in + let typ = EConstr.of_constr typ in print_typed_value_in_env env sigma (trm, typ) (*********************) @@ -762,7 +763,9 @@ let print_opaque_name qid = | IndRef (sp,_) -> print_inductive sp | ConstructRef cstr as gr -> + let open EConstr in let ty = Universes.unsafe_type_of_global gr in + let ty = EConstr.of_constr ty in print_typed_value (mkConstruct cstr, ty) | VarRef id -> env |> lookup_named id |> NamedDecl.set_id id |> print_named_decl diff --git a/printing/prettyp.mli b/printing/prettyp.mli index 0eab15579..38e111034 100644 --- a/printing/prettyp.mli +++ b/printing/prettyp.mli @@ -27,11 +27,11 @@ val print_full_context_typ : unit -> std_ppcmds val print_full_pure_context : unit -> std_ppcmds val print_sec_context : reference -> std_ppcmds val print_sec_context_typ : reference -> std_ppcmds -val print_judgment : env -> Evd.evar_map -> unsafe_judgment -> std_ppcmds +val print_judgment : env -> Evd.evar_map -> EConstr.unsafe_judgment -> std_ppcmds val print_safe_judgment : env -> Evd.evar_map -> Safe_typing.judgment -> std_ppcmds val print_eval : reduction_function -> env -> Evd.evar_map -> - Constrexpr.constr_expr -> unsafe_judgment -> std_ppcmds + Constrexpr.constr_expr -> EConstr.unsafe_judgment -> std_ppcmds val print_name : reference or_by_notation -> std_ppcmds val print_opaque_name : reference -> std_ppcmds @@ -69,8 +69,8 @@ type object_pr = { print_named_decl : Context.Named.Declaration.t -> std_ppcmds; print_library_entry : bool -> (object_name * Lib.node) -> std_ppcmds option; print_context : bool -> int option -> Lib.library_segment -> std_ppcmds; - print_typed_value_in_env : Environ.env -> Evd.evar_map -> Term.constr * Term.types -> Pp.std_ppcmds; - print_eval : reduction_function -> env -> Evd.evar_map -> Constrexpr.constr_expr -> unsafe_judgment -> std_ppcmds + print_typed_value_in_env : Environ.env -> Evd.evar_map -> EConstr.constr * EConstr.types -> Pp.std_ppcmds; + print_eval : reduction_function -> env -> Evd.evar_map -> Constrexpr.constr_expr -> EConstr.unsafe_judgment -> std_ppcmds } val set_object_pr : object_pr -> unit diff --git a/printing/printer.ml b/printing/printer.ml index 4a6c83bd7..ba4b68296 100644 --- a/printing/printer.ml +++ b/printing/printer.ml @@ -99,7 +99,6 @@ let pr_type_core goal_concl_style env sigma t = let pr_ltype_core goal_concl_style env sigma t = pr_lconstr_expr (extern_type goal_concl_style env sigma t) -let pr_goal_concl_style_env env = pr_ltype_core true env let pr_ltype_env env = pr_ltype_core false env let pr_type_env env = pr_type_core false env @@ -110,8 +109,13 @@ let pr_type t = let (sigma, env) = get_current_context () in pr_type_env env sigma t +let pr_etype_env env sigma c = pr_type_env env sigma (EConstr.to_constr sigma c) +let pr_letype_env env sigma c = pr_ltype_env env sigma (EConstr.to_constr sigma c) +let pr_goal_concl_style_env env sigma c = + pr_ltype_core true env sigma (EConstr.to_constr sigma c) + let pr_ljudge_env env sigma j = - (pr_lconstr_env env sigma j.uj_val, pr_lconstr_env env sigma j.uj_type) + (pr_leconstr_env env sigma j.uj_val, pr_leconstr_env env sigma j.uj_type) let pr_ljudge j = let (sigma, env) = get_current_context () in @@ -390,7 +394,7 @@ let pr_transparent_state (ids, csts) = let default_pr_goal gs = let (g,sigma) = Goal.V82.nf_evar (project gs) (sig_it gs) in let env = Goal.V82.env sigma g in - let concl = EConstr.Unsafe.to_constr (Goal.V82.concl sigma g) in + let concl = Goal.V82.concl sigma g in let goal = pr_context_of env sigma ++ cut () ++ str "============================" ++ cut () ++ @@ -413,7 +417,7 @@ let pr_goal_name sigma g = let pr_concl n sigma g = let (g,sigma) = Goal.V82.nf_evar sigma g in let env = Goal.V82.env sigma g in - let pc = pr_goal_concl_style_env env sigma (EConstr.Unsafe.to_constr (Goal.V82.concl sigma g)) in + let pc = pr_goal_concl_style_env env sigma (Goal.V82.concl sigma g) in str (emacs_str "") ++ str "subgoal " ++ int n ++ pr_goal_tag g ++ pr_goal_name sigma g ++ str " is:" ++ cut () ++ str" " ++ pc diff --git a/printing/printer.mli b/printing/printer.mli index 7521468e2..504392e35 100644 --- a/printing/printer.mli +++ b/printing/printer.mli @@ -43,6 +43,9 @@ val pr_econstr : EConstr.t -> std_ppcmds val pr_leconstr_env : env -> evar_map -> EConstr.t -> std_ppcmds val pr_leconstr : EConstr.t -> std_ppcmds +val pr_etype_env : env -> evar_map -> EConstr.types -> std_ppcmds +val pr_letype_env : env -> evar_map -> EConstr.types -> std_ppcmds + val pr_open_constr_env : env -> evar_map -> open_constr -> std_ppcmds val pr_open_constr : open_constr -> std_ppcmds @@ -55,7 +58,7 @@ val pr_constr_under_binders : constr_under_binders -> std_ppcmds val pr_lconstr_under_binders_env : env -> evar_map -> constr_under_binders -> std_ppcmds val pr_lconstr_under_binders : constr_under_binders -> std_ppcmds -val pr_goal_concl_style_env : env -> evar_map -> types -> std_ppcmds +val pr_goal_concl_style_env : env -> evar_map -> EConstr.types -> std_ppcmds val pr_ltype_env : env -> evar_map -> types -> std_ppcmds val pr_ltype : types -> std_ppcmds @@ -65,8 +68,8 @@ val pr_type : types -> std_ppcmds val pr_closed_glob_env : env -> evar_map -> closed_glob_constr -> std_ppcmds val pr_closed_glob : closed_glob_constr -> std_ppcmds -val pr_ljudge_env : env -> evar_map -> unsafe_judgment -> std_ppcmds * std_ppcmds -val pr_ljudge : unsafe_judgment -> std_ppcmds * std_ppcmds +val pr_ljudge_env : env -> evar_map -> EConstr.unsafe_judgment -> std_ppcmds * std_ppcmds +val pr_ljudge : EConstr.unsafe_judgment -> std_ppcmds * std_ppcmds val pr_lglob_constr_env : env -> glob_constr -> std_ppcmds val pr_lglob_constr : glob_constr -> std_ppcmds diff --git a/tactics/class_tactics.ml b/tactics/class_tactics.ml index 2f8af6b44..84ca0aa8f 100644 --- a/tactics/class_tactics.ml +++ b/tactics/class_tactics.ml @@ -185,8 +185,7 @@ let set_typeclasses_depth = optwrite = set_typeclasses_depth; } let pr_ev evs ev = - Printer.pr_constr_env (Goal.V82.env evs ev) evs - (Evarutil.nf_evar evs (EConstr.Unsafe.to_constr (Goal.V82.concl evs ev))) + Printer.pr_econstr_env (Goal.V82.env evs ev) evs (Goal.V82.concl evs ev) (** Typeclasses instance search tactic / eauto *) @@ -764,7 +763,7 @@ module V85 = struct if foundone == None && !typeclasses_debug > 0 then Feedback.msg_debug (pr_depth info.auto_depth ++ str": no match for " ++ - Printer.pr_constr_env (Goal.V82.env s gl) s (EConstr.Unsafe.to_constr concl) ++ + Printer.pr_econstr_env (Goal.V82.env s gl) s concl ++ spc () ++ str ", " ++ int (List.length poss) ++ str" possibilities"); match foundone with @@ -1005,7 +1004,7 @@ module Search = struct if !typeclasses_debug > 0 then Feedback.msg_debug (pr_depth info.search_depth ++ str": looking for " ++ - Printer.pr_constr_env (Goal.env gl) s (EConstr.Unsafe.to_constr concl) ++ + Printer.pr_econstr_env (Goal.env gl) s concl ++ (if backtrack then str" with backtracking" else str" without backtracking")); let secvars = compute_secvars gl in @@ -1120,7 +1119,7 @@ module Search = struct if !foundone == false && !typeclasses_debug > 0 then Feedback.msg_debug (pr_depth info.search_depth ++ str": no match for " ++ - Printer.pr_constr_env (Goal.env gl) s (EConstr.Unsafe.to_constr concl) ++ + Printer.pr_econstr_env (Goal.env gl) s concl ++ spc () ++ str ", " ++ int (List.length poss) ++ str" possibilities"); match e with diff --git a/tactics/hints.ml b/tactics/hints.ml index 9e9635e8a..2446b6996 100644 --- a/tactics/hints.ml +++ b/tactics/hints.ml @@ -792,7 +792,7 @@ let make_apply_entry env sigma (eapply,hnf,verbose) pri poly ?(name=PathAny) (c, else begin if not eapply then failwith "make_apply_entry"; if verbose then - Feedback.msg_info (str "the hint: eapply " ++ pr_lconstr (EConstr.Unsafe.to_constr c) ++ + Feedback.msg_info (str "the hint: eapply " ++ pr_leconstr_env env sigma' c ++ str " will only be used by eauto"); (Some hd, { pri = (match pri with None -> nb_hyp sigma' cty + nmiss | Some p -> p); @@ -813,7 +813,7 @@ let pr_hint_term env sigma ctx = function | IsGlobRef gr -> pr_global gr | IsConstr (c, ctx) -> let sigma = Evd.merge_context_set Evd.univ_flexible sigma ctx in - pr_constr_env env sigma (EConstr.Unsafe.to_constr c) + pr_econstr_env env sigma c (** We need an object to record the side-effect of registering global universes associated with a hint. *) @@ -863,7 +863,7 @@ let make_resolves env sigma flags pri poly ?name cr = in if List.is_empty ents then user_err ~hdr:"Hint" - (pr_lconstr (EConstr.Unsafe.to_constr c) ++ spc() ++ + (pr_leconstr_env env sigma c ++ spc() ++ (if pi1 flags then str"cannot be used as a hint." else str "can be used as a hint only for eauto.")); ents @@ -1360,7 +1360,7 @@ let make_db_list dbnames = (* Functions for printing the hints *) (**************************************************************************) -let pr_hint_elt (c, _, _) = pr_constr (EConstr.Unsafe.to_constr c) +let pr_hint_elt (c, _, _) = pr_econstr c let pr_hint h = match h.obj with | Res_pf (c, _) -> (str"simple apply " ++ pr_hint_elt c) diff --git a/tactics/leminv.ml b/tactics/leminv.ml index ef3bfc9d0..2d59285e6 100644 --- a/tactics/leminv.ml +++ b/tactics/leminv.ml @@ -42,7 +42,7 @@ let nlocal_def (na, b, t) = let no_inductive_inconstr env sigma constr = (str "Cannot recognize an inductive predicate in " ++ - pr_lconstr_env env sigma (EConstr.Unsafe.to_constr constr) ++ + pr_leconstr_env env sigma constr ++ str "." ++ spc () ++ str "If there is one, may be the structure of the arity" ++ spc () ++ str "or of the type of constructors" ++ spc () ++ str "is hidden by constant definitions.") @@ -277,14 +277,12 @@ let lemInv id c gls = Proofview.V82.of_tactic (Clenvtac.res_pf clause ~flags:(Unification.elim_flags ()) ~with_evars:false) gls with | NoSuchBinding -> - let c = EConstr.Unsafe.to_constr c in user_err - (hov 0 (pr_constr c ++ spc () ++ str "does not refer to an inversion lemma.")) + (hov 0 (pr_econstr_env (Refiner.pf_env gls) (Refiner.project gls) c ++ spc () ++ str "does not refer to an inversion lemma.")) | UserError (a,b) -> - let c = EConstr.Unsafe.to_constr c in user_err ~hdr:"LemInv" (str "Cannot refine current goal with the lemma " ++ - pr_lconstr_env (Refiner.pf_env gls) (Refiner.project gls) c) + pr_leconstr_env (Refiner.pf_env gls) (Refiner.project gls) c) let lemInv_gen id c = try_intros_until (fun id -> Proofview.V82.tactic (lemInv id c)) id diff --git a/tactics/tactics.ml b/tactics/tactics.ml index 4e833eb55..dcaa15fd8 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -3073,7 +3073,7 @@ let warn_unused_intro_pattern = strbrk"Unused introduction " ++ str (String.plural (List.length names) "pattern") ++ str": " ++ prlist_with_sep spc (Miscprint.pr_intro_pattern - (fun c -> Printer.pr_constr (EConstr.Unsafe.to_constr (fst (run_delayed (Global.env()) Evd.empty c))))) names) + (fun c -> Printer.pr_econstr (fst (run_delayed (Global.env()) Evd.empty c)))) names) let check_unused_names names = if not (List.is_empty names) && Flags.is_verbose () then diff --git a/toplevel/himsg.ml b/toplevel/himsg.ml index 7eb189ef5..3c2fe46b3 100644 --- a/toplevel/himsg.ml +++ b/toplevel/himsg.ml @@ -86,11 +86,11 @@ let rec contract3' env a b c = function let to_unsafe_judgment j = on_judgment EConstr.Unsafe.to_constr j let j_nf_betaiotaevar sigma j = - { uj_val = Evarutil.nf_evar sigma j.uj_val; - uj_type = EConstr.Unsafe.to_constr (Reductionops.nf_betaiota sigma (EConstr.of_constr j.uj_type)) } + { uj_val = Evarutil.nf_evar sigma (EConstr.Unsafe.to_constr j.uj_val); + uj_type = EConstr.Unsafe.to_constr (Reductionops.nf_betaiota sigma j.uj_type) } let jv_nf_betaiotaevar sigma jl = - Array.map (j_nf_betaiotaevar sigma) jl + Array.map (fun j -> on_judgment EConstr.of_constr (j_nf_betaiotaevar sigma j)) jl (** Printers *) @@ -172,7 +172,6 @@ let explain_unbound_var env v = let explain_not_type env sigma j = let j = Evarutil.j_nf_evar sigma j in - let j = to_unsafe_judgment j in let pe = pr_ne_context_of (str "In environment") env sigma in let pc,pt = pr_ljudge_env env sigma j in pe ++ str "The term" ++ brk(1,1) ++ pc ++ spc () ++ @@ -180,17 +179,15 @@ let explain_not_type env sigma j = str "which should be Set, Prop or Type." let explain_bad_assumption env sigma j = - let j = to_unsafe_judgment j in let pe = pr_ne_context_of (str "In environment") env sigma in let pc,pt = pr_ljudge_env env sigma j in pe ++ str "Cannot declare a variable or hypothesis over the term" ++ brk(1,1) ++ pc ++ spc () ++ str "of type" ++ spc () ++ pt ++ spc () ++ str "because this term is not a type." -let explain_reference_variables id c = - let c = EConstr.Unsafe.to_constr c in +let explain_reference_variables sigma id c = (* c is intended to be a global reference *) - let pc = pr_global (Globnames.global_of_constr c) in + let pc = pr_global (fst (Termops.global_of_constr sigma c)) in pc ++ strbrk " depends on the variable " ++ pr_id id ++ strbrk " which is not declared in the context." @@ -207,11 +204,10 @@ let pr_puniverses f env (c,u) = else mt()) let explain_elim_arity env sigma ind sorts c pj okinds = - let c = EConstr.Unsafe.to_constr c in let pj = to_unsafe_judgment pj in let env = make_all_name_different env in let pi = pr_inductive env (fst ind) in - let pc = pr_lconstr_env env sigma c in + let pc = pr_leconstr_env env sigma c in let msg = match okinds with | Some(kp,ki,explanation) -> let pki = pr_sort_family ki in @@ -267,13 +263,11 @@ let explain_number_branches env sigma cj expn = str "expects " ++ int expn ++ str " branches." let explain_ill_formed_branch env sigma c ci actty expty = - let c = EConstr.Unsafe.to_constr c in let actty = EConstr.Unsafe.to_constr actty in let expty = EConstr.Unsafe.to_constr expty in let simp t = Reduction.nf_betaiota env (Evarutil.nf_evar sigma t) in - let c = Evarutil.nf_evar sigma c in let env = make_all_name_different env in - let pc = pr_lconstr_env env sigma c in + let pc = pr_leconstr_env env sigma c in let pa, pe = pr_explicit env sigma (simp actty) (simp expty) in strbrk "In pattern-matching on term" ++ brk(1,1) ++ pc ++ spc () ++ strbrk "the branch for constructor" ++ spc () ++ @@ -282,11 +276,9 @@ let explain_ill_formed_branch env sigma c ci actty expty = str "which should be" ++ brk(1,1) ++ pe ++ str "." let explain_generalization env sigma (name,var) j = - let var = EConstr.Unsafe.to_constr var in - let j = to_unsafe_judgment j in let pe = pr_ne_context_of (str "In environment") env sigma in - let pv = pr_ltype_env env sigma var in - let (pc,pt) = pr_ljudge_env (push_rel_assum (name,EConstr.of_constr var) env) sigma j in + let pv = pr_letype_env env sigma var in + let (pc,pt) = pr_ljudge_env (push_rel_assum (name,var) env) sigma j in pe ++ str "Cannot generalize" ++ brk(1,1) ++ pv ++ spc () ++ str "over" ++ brk(1,1) ++ pc ++ str "," ++ spc () ++ str "it has type" ++ spc () ++ pt ++ @@ -297,19 +289,16 @@ let explain_unification_error env sigma p1 p2 = function | Some e -> let rec aux p1 p2 = function | OccurCheck (evk,rhs) -> - let rhs = EConstr.to_constr sigma rhs in [str "cannot define " ++ quote (pr_existential_key sigma evk) ++ - strbrk " with term " ++ pr_lconstr_env env sigma rhs ++ + strbrk " with term " ++ pr_leconstr_env env sigma rhs ++ strbrk " that would depend on itself"] | NotClean ((evk,args),env,c) -> - let c = EConstr.to_constr sigma c in - let args = Array.map (EConstr.to_constr sigma) args in [str "cannot instantiate " ++ quote (pr_existential_key sigma evk) - ++ strbrk " because " ++ pr_lconstr_env env sigma c ++ + ++ strbrk " because " ++ pr_leconstr_env env sigma c ++ strbrk " is not in its scope" ++ (if Array.is_empty args then mt() else strbrk ": available arguments are " ++ - pr_sequence (pr_lconstr_env env sigma) (List.rev (Array.to_list args)))] + pr_sequence (pr_leconstr_env env sigma) (List.rev (Array.to_list args)))] | NotSameArgSize | NotSameHead | NoCanonicalStructure -> (* Error speaks from itself *) [] | ConversionFailed (env,t1,t2) -> @@ -357,11 +346,9 @@ let explain_unification_error env sigma p1 p2 = function prlist_with_sep pr_semicolon (fun x -> x) l ++ str ")" let explain_actual_type env sigma j t reason = - let j = to_unsafe_judgment j in - let t = EConstr.Unsafe.to_constr t in let env = make_all_name_different env in let j = j_nf_betaiotaevar sigma j in - let t = Reductionops.nf_betaiota sigma (EConstr.of_constr t) in + let t = Reductionops.nf_betaiota sigma t in let t = EConstr.Unsafe.to_constr t in (** Actually print *) let pe = pr_ne_context_of (str "In environment") env sigma in @@ -377,14 +364,11 @@ let explain_actual_type env sigma j t reason = let explain_cant_apply_bad_type env sigma (n,exptyp,actualtyp) rator randl = let exptyp = EConstr.Unsafe.to_constr exptyp in - let actualtyp = EConstr.Unsafe.to_constr actualtyp in - let randl = Array.map to_unsafe_judgment randl in let randl = jv_nf_betaiotaevar sigma randl in let exptyp = Evarutil.nf_evar sigma exptyp in - let actualtyp = Reductionops.nf_betaiota sigma (EConstr.of_constr actualtyp) in + let actualtyp = Reductionops.nf_betaiota sigma actualtyp in let actualtyp = EConstr.Unsafe.to_constr actualtyp in let rator = Evarutil.j_nf_evar sigma rator in - let rator = to_unsafe_judgment rator in let env = make_all_name_different env in let actualtyp, exptyp = pr_explicit env sigma actualtyp exptyp in let nargs = Array.length randl in @@ -448,7 +432,7 @@ let explain_not_product env sigma c = (* TODO: use the names *) (* (co)fixpoints *) let explain_ill_formed_rec_body env sigma err names i fixenv vdefj = - let pr_lconstr_env env sigma c = pr_lconstr_env env sigma (EConstr.Unsafe.to_constr c) in + let pr_lconstr_env env sigma c = pr_leconstr_env env sigma c in let prt_name i = match names.(i) with Name id -> str "Recursive definition of " ++ pr_id id @@ -549,17 +533,14 @@ let explain_ill_typed_rec_body env sigma i names vdefj vargs = str "while it should be" ++ spc () ++ pv ++ str "." let explain_cant_find_case_type env sigma c = - let c = EConstr.Unsafe.to_constr c in - let c = Evarutil.nf_evar sigma c in let env = make_all_name_different env in - let pe = pr_lconstr_env env sigma c in + let pe = pr_leconstr_env env sigma c in str "Cannot infer the return type of pattern-matching on" ++ ws 1 ++ pe ++ str "." let explain_occur_check env sigma ev rhs = - let rhs = EConstr.to_constr sigma rhs in let env = make_all_name_different env in - let pt = pr_lconstr_env env sigma rhs in + let pt = pr_leconstr_env env sigma rhs in str "Cannot define " ++ pr_existential_key sigma ev ++ str " with term" ++ brk(1,1) ++ pt ++ spc () ++ str "that would depend on itself." @@ -665,12 +646,9 @@ let explain_cannot_unify env sigma m n e = str "with" ++ brk(1,1) ++ pn ++ ppreason ++ str "." let explain_cannot_unify_local env sigma m n subn = - let m = EConstr.to_constr sigma m in - let n = EConstr.to_constr sigma n in - let subn = EConstr.to_constr sigma subn in - let pm = pr_lconstr_env env sigma m in - let pn = pr_lconstr_env env sigma n in - let psubn = pr_lconstr_env env sigma subn in + let pm = pr_leconstr_env env sigma m in + let pn = pr_leconstr_env env sigma n in + let psubn = pr_leconstr_env env sigma subn in str "Unable to unify" ++ brk(1,1) ++ pm ++ spc () ++ str "with" ++ brk(1,1) ++ pn ++ spc () ++ str "as" ++ brk(1,1) ++ psubn ++ str " contains local variables." @@ -740,7 +718,7 @@ let explain_type_error env sigma err = | BadAssumption c -> explain_bad_assumption env sigma c | ReferenceVariables (id,c) -> - explain_reference_variables id c + explain_reference_variables sigma id c | ElimArity (ind, aritylst, c, pj, okinds) -> explain_elim_arity env sigma ind aritylst c pj okinds | CaseNotInductive cj -> @@ -1117,8 +1095,7 @@ let explain_non_linear_proof c = spc () ++ str "because a metavariable has several occurrences." let explain_meta_in_type c = - let c = EConstr.Unsafe.to_constr c in - str "In refiner, a meta appears in the type " ++ brk(1,1) ++ pr_lconstr c ++ + str "In refiner, a meta appears in the type " ++ brk(1,1) ++ pr_leconstr c ++ str " of another meta" let explain_no_such_hyp id = @@ -1170,7 +1147,7 @@ let error_ill_formed_constructor env id c v nparams nargs = let pr_ltype_using_barendregt_convention_env env c = (* Use goal_concl_style as an approximation of Barendregt's convention (?) *) - quote (pr_goal_concl_style_env env Evd.empty c) + quote (pr_goal_concl_style_env env Evd.empty (EConstr.of_constr c)) let error_bad_ind_parameters env c n v1 v2 = let pc = pr_ltype_using_barendregt_convention_env env c in @@ -1364,7 +1341,6 @@ let map_ptype_error f = function let explain_reduction_tactic_error = function | Tacred.InvalidAbstraction (env,sigma,c,(env',e)) -> let e = map_ptype_error EConstr.of_constr e in - let c = EConstr.to_constr sigma c in str "The abstracted term" ++ spc () ++ quote (pr_goal_concl_style_env env sigma c) ++ spc () ++ str "is not well typed." ++ fnl () ++ diff --git a/toplevel/vernacentries.ml b/toplevel/vernacentries.ml index 25b639fb0..ee2aacfc5 100644 --- a/toplevel/vernacentries.ml +++ b/toplevel/vernacentries.ml @@ -1580,15 +1580,16 @@ let vernac_check_may_eval redexp glopt rc = let c = nf c in let j = if Evarutil.has_undefined_evars sigma' (EConstr.of_constr c) then - let j = Evarutil.j_nf_evar sigma' (Retyping.get_judgment_of env sigma' (EConstr.of_constr c)) in - Termops.on_judgment EConstr.Unsafe.to_constr j + Evarutil.j_nf_evar sigma' (Retyping.get_judgment_of env sigma' (EConstr.of_constr c)) else (* OK to call kernel which does not support evars *) - Arguments_renaming.rename_typing env c in + Termops.on_judgment EConstr.of_constr (Arguments_renaming.rename_typing env c) + in match redexp with | None -> - let l = Evar.Set.union (Evd.evars_of_term j.Environ.uj_val) (Evd.evars_of_term j.Environ.uj_type) in - let j = { j with Environ.uj_type = EConstr.Unsafe.to_constr (Reductionops.nf_betaiota sigma' (EConstr.of_constr j.Environ.uj_type)) } in + let evars_of_term c = Evarutil.undefined_evars_of_term sigma' c in + let l = Evar.Set.union (evars_of_term j.Environ.uj_val) (evars_of_term j.Environ.uj_type) in + let j = { j with Environ.uj_type = Reductionops.nf_betaiota sigma' j.Environ.uj_type } in Feedback.msg_notice (print_judgment env sigma' j ++ pr_ne_evar_set (fnl () ++ str "where" ++ fnl ()) (mt ()) sigma' l ++ Printer.pr_universe_ctx sigma uctx) -- cgit v1.2.3 From a5499688bd76def8de3d8e1089a49c7a08430903 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 25 Nov 2016 14:54:40 +0100 Subject: Funind API using EConstr. --- plugins/funind/functional_principles_proofs.ml | 364 ++++++++++++------------ plugins/funind/functional_principles_proofs.mli | 6 +- plugins/funind/functional_principles_types.ml | 6 +- plugins/funind/functional_principles_types.mli | 2 +- plugins/funind/g_indfun.ml4 | 3 - plugins/funind/glob_term_to_relation.ml | 2 +- plugins/funind/indfun.ml | 45 +-- plugins/funind/indfun.mli | 4 +- plugins/funind/indfun_common.ml | 56 +++- plugins/funind/indfun_common.mli | 23 +- plugins/funind/invfun.ml | 212 +++++++------- plugins/funind/recdef.ml | 342 ++++++++++++---------- plugins/funind/recdef.mli | 2 +- proofs/goal.ml | 5 +- proofs/goal.mli | 2 +- 15 files changed, 579 insertions(+), 495 deletions(-) diff --git a/plugins/funind/functional_principles_proofs.ml b/plugins/funind/functional_principles_proofs.ml index c98cdc467..656924e38 100644 --- a/plugins/funind/functional_principles_proofs.ml +++ b/plugins/funind/functional_principles_proofs.ml @@ -2,6 +2,7 @@ open Printer open CErrors open Util open Term +open EConstr open Vars open Namegen open Names @@ -18,6 +19,12 @@ open Context.Rel.Declaration module RelDecl = Context.Rel.Declaration +let local_assum (na, t) = + RelDecl.LocalAssum (na, EConstr.Unsafe.to_constr t) + +let local_def (na, b, t) = + RelDecl.LocalDef (na, EConstr.Unsafe.to_constr b, EConstr.Unsafe.to_constr t) + (* let msgnl = Pp.msgnl *) (* @@ -132,16 +139,16 @@ let refine c = let thin l = Proofview.V82.of_tactic (Tactics.clear l) -let eq_constr u v = eq_constr_nounivs u v +let eq_constr sigma u v = EConstr.eq_constr_nounivs sigma u v -let is_trivial_eq t = +let is_trivial_eq sigma t = let res = try begin - match kind_of_term t with - | App(f,[|_;t1;t2|]) when eq_constr f (Lazy.force eq) -> - eq_constr t1 t2 - | App(f,[|t1;a1;t2;a2|]) when eq_constr f (jmeq ()) -> - eq_constr t1 t2 && eq_constr a1 a2 + match EConstr.kind sigma t with + | App(f,[|_;t1;t2|]) when eq_constr sigma f (Lazy.force eq) -> + eq_constr sigma t1 t2 + | App(f,[|t1;a1;t2;a2|]) when eq_constr sigma f (jmeq ()) -> + eq_constr sigma t1 t2 && eq_constr sigma a1 a2 | _ -> false end with e when CErrors.noncritical e -> false @@ -149,34 +156,33 @@ let is_trivial_eq t = (* observe (str "is_trivial_eq " ++ Printer.pr_lconstr t ++ (if res then str " true" else str " false")); *) res -let rec incompatible_constructor_terms t1 t2 = - let c1,arg1 = decompose_app t1 - and c2,arg2 = decompose_app t2 +let rec incompatible_constructor_terms sigma t1 t2 = + let c1,arg1 = decompose_app sigma t1 + and c2,arg2 = decompose_app sigma t2 in - (not (eq_constr t1 t2)) && - isConstruct c1 && isConstruct c2 && + (not (eq_constr sigma t1 t2)) && + isConstruct sigma c1 && isConstruct sigma c2 && ( - not (eq_constr c1 c2) || - List.exists2 incompatible_constructor_terms arg1 arg2 + not (eq_constr sigma c1 c2) || + List.exists2 (incompatible_constructor_terms sigma) arg1 arg2 ) -let is_incompatible_eq t = +let is_incompatible_eq sigma t = let res = try - match kind_of_term t with - | App(f,[|_;t1;t2|]) when eq_constr f (Lazy.force eq) -> - incompatible_constructor_terms t1 t2 - | App(f,[|u1;t1;u2;t2|]) when eq_constr f (jmeq ()) -> - (eq_constr u1 u2 && - incompatible_constructor_terms t1 t2) + match EConstr.kind sigma t with + | App(f,[|_;t1;t2|]) when eq_constr sigma f (Lazy.force eq) -> + incompatible_constructor_terms sigma t1 t2 + | App(f,[|u1;t1;u2;t2|]) when eq_constr sigma f (jmeq ()) -> + (eq_constr sigma u1 u2 && + incompatible_constructor_terms sigma t1 t2) | _ -> false with e when CErrors.noncritical e -> false in - if res then observe (str "is_incompatible_eq " ++ Printer.pr_lconstr t); + if res then observe (str "is_incompatible_eq " ++ Printer.pr_leconstr t); res let change_hyp_with_using msg hyp_id t tac : tactic = - let t = EConstr.of_constr t in fun g -> let prov_id = pf_get_new_id hyp_id g in tclTHENS @@ -204,47 +210,44 @@ let prove_trivial_eq h_id context (constructor,type_of_term,term) = (List.map mkVar context_hyps) in let to_refine = applist(mkVar h_id,List.rev context_hyps') in - let to_refine = EConstr.of_constr to_refine in refine to_refine g ) ] -let find_rectype env c = - let (t, l) = decompose_app (Reduction.whd_betaiotazeta env c) in - match kind_of_term t with +let find_rectype env sigma c = + let (t, l) = decompose_app sigma (Reductionops.whd_betaiotazeta sigma c) in + match EConstr.kind sigma t with | Ind ind -> (t, l) | Construct _ -> (t,l) | _ -> raise Not_found -let isAppConstruct ?(env=Global.env ()) t = +let isAppConstruct ?(env=Global.env ()) sigma t = try - let t',l = find_rectype (Global.env ()) t in - observe (str "isAppConstruct : " ++ Printer.pr_lconstr t ++ str " -> " ++ Printer.pr_lconstr (applist (t',l))); + let t',l = find_rectype env sigma t in + observe (str "isAppConstruct : " ++ Printer.pr_leconstr t ++ str " -> " ++ Printer.pr_leconstr (applist (t',l))); true with Not_found -> false let nf_betaiotazeta = (* Reductionops.local_strong Reductionops.whd_betaiotazeta *) - let clos_norm_flags flgs env sigma t = - CClosure.norm_val (CClosure.create_clos_infos flgs env) (CClosure.inject (Reductionops.nf_evar sigma t)) in - clos_norm_flags CClosure.betaiotazeta Environ.empty_env Evd.empty + Reductionops.clos_norm_flags CClosure.betaiotazeta Environ.empty_env Evd.empty let change_eq env sigma hyp_id (context:Context.Rel.t) x t end_of_type = let nochange ?t' msg = begin - observe (str ("Not treating ( "^msg^" )") ++ pr_lconstr t ++ str " " ++ match t' with None -> str "" | Some t -> Printer.pr_lconstr t ); + observe (str ("Not treating ( "^msg^" )") ++ pr_leconstr t ++ str " " ++ match t' with None -> str "" | Some t -> Printer.pr_leconstr t ); failwith "NoChange"; end in - let eq_constr c1 c2 = Evarconv.e_conv env (ref sigma) (EConstr.of_constr c1) (EConstr.of_constr c2) in - if not (noccurn 1 end_of_type) + let eq_constr c1 c2 = Evarconv.e_conv env (ref sigma) c1 c2 in + if not (noccurn sigma 1 end_of_type) then nochange "dependent"; (* if end_of_type depends on this term we don't touch it *) - if not (isApp t) then nochange "not an equality"; - let f_eq,args = destApp t in + if not (isApp sigma t) then nochange "not an equality"; + let f_eq,args = destApp sigma t in let constructor,t1,t2,t1_typ = try if (eq_constr f_eq (Lazy.force eq)) @@ -261,32 +264,32 @@ let change_eq env sigma hyp_id (context:Context.Rel.t) x t end_of_type = else nochange "not an equality" with e when CErrors.noncritical e -> nochange "not an equality" in - if not ((closed0 (fst t1)) && (closed0 (snd t1)))then nochange "not a closed lhs"; + if not ((closed0 sigma (fst t1)) && (closed0 sigma (snd t1)))then nochange "not a closed lhs"; let rec compute_substitution sub t1 t2 = (* observe (str "compute_substitution : " ++ pr_lconstr t1 ++ str " === " ++ pr_lconstr t2); *) - if isRel t2 + if isRel sigma t2 then - let t2 = destRel t2 in + let t2 = destRel sigma t2 in begin try let t1' = Int.Map.find t2 sub in if not (eq_constr t1 t1') then nochange "twice bound variable"; sub with Not_found -> - assert (closed0 t1); + assert (closed0 sigma t1); Int.Map.add t2 t1 sub end - else if isAppConstruct t1 && isAppConstruct t2 + else if isAppConstruct sigma t1 && isAppConstruct sigma t2 then begin - let c1,args1 = find_rectype env t1 - and c2,args2 = find_rectype env t2 + let c1,args1 = find_rectype env sigma t1 + and c2,args2 = find_rectype env sigma t2 in if not (eq_constr c1 c2) then nochange "cannot solve (diff)"; List.fold_left2 compute_substitution sub args1 args2 end else - if (eq_constr t1 t2) then sub else nochange ~t':(make_refl_eq constructor (Reduction.whd_all env t1) t2) "cannot solve (diff)" + if (eq_constr t1 t2) then sub else nochange ~t':(make_refl_eq constructor (Reductionops.whd_all env sigma t1) t2) "cannot solve (diff)" in let sub = compute_substitution Int.Map.empty (snd t1) (snd t2) in let sub = compute_substitution sub (fst t1) (fst t2) in @@ -312,19 +315,18 @@ let change_eq env sigma hyp_id (context:Context.Rel.t) x t end_of_type = try let witness = Int.Map.find i sub in if is_local_def decl then anomaly (Pp.str "can not redefine a rel!"); - (pop end_of_type,ctxt_size,mkLetIn (RelDecl.get_name decl, witness, RelDecl.get_type decl, witness_fun)) + (pop end_of_type,ctxt_size,mkLetIn (RelDecl.get_name decl, witness, EConstr.of_constr (RelDecl.get_type decl), witness_fun)) with Not_found -> - (Term.mkProd_or_LetIn decl end_of_type, ctxt_size + 1, mkLambda_or_LetIn decl witness_fun) + (mkProd_or_LetIn decl end_of_type, ctxt_size + 1, mkLambda_or_LetIn decl witness_fun) ) 1 (new_end_of_type,0,witness_fun) context in let new_type_of_hyp = - Reductionops.nf_betaiota Evd.empty (EConstr.of_constr new_type_of_hyp) in - let new_type_of_hyp = EConstr.Unsafe.to_constr new_type_of_hyp in + Reductionops.nf_betaiota sigma new_type_of_hyp in let new_ctxt,new_end_of_type = - decompose_prod_n_assum ctxt_size new_type_of_hyp + decompose_prod_n_assum sigma ctxt_size new_type_of_hyp in let prove_new_hyp : tactic = tclTHEN @@ -333,7 +335,6 @@ let change_eq env sigma hyp_id (context:Context.Rel.t) x t end_of_type = let all_ids = pf_ids_of_hyps g in let new_ids,_ = list_chop ctxt_size all_ids in let to_refine = applist(witness_fun,List.rev_map mkVar new_ids) in - let to_refine = EConstr.of_constr to_refine in let evm, _ = pf_apply Typing.type_of g to_refine in tclTHEN (Refiner.tclEVARS evm) (refine to_refine) g ) @@ -358,21 +359,21 @@ let change_eq env sigma hyp_id (context:Context.Rel.t) x t end_of_type = new_ctxt,new_end_of_type,simpl_eq_tac -let is_property (ptes_info:ptes_info) t_x full_type_of_hyp = - if isApp t_x +let is_property sigma (ptes_info:ptes_info) t_x full_type_of_hyp = + if isApp sigma t_x then - let pte,args = destApp t_x in - if isVar pte && Array.for_all closed0 args + let pte,args = destApp sigma t_x in + if isVar sigma pte && Array.for_all (closed0 sigma) args then try - let info = Id.Map.find (destVar pte) ptes_info in + let info = Id.Map.find (destVar sigma pte) ptes_info in info.is_valid full_type_of_hyp with Not_found -> false else false else false -let isLetIn t = - match kind_of_term t with +let isLetIn sigma t = + match EConstr.kind sigma t with | LetIn _ -> true | _ -> false @@ -392,8 +393,9 @@ let rewrite_until_var arg_num eq_ids : tactic = will break the Guard when trying to save the Lemma. *) let test_var g = - let _,args = destApp (pf_concl g) in - not ((isConstruct args.(arg_num)) || isAppConstruct args.(arg_num)) + let sigma = project g in + let _,args = destApp sigma (EConstr.of_constr (pf_concl g)) in + not ((isConstruct sigma args.(arg_num)) || isAppConstruct sigma args.(arg_num)) in let rec do_rewrite eq_ids g = if test_var g @@ -403,7 +405,7 @@ let rewrite_until_var arg_num eq_ids : tactic = | [] -> anomaly (Pp.str "Cannot find a way to prove recursive property"); | eq_id::eq_ids -> tclTHEN - (tclTRY (Proofview.V82.of_tactic (Equality.rewriteRL (EConstr.mkVar eq_id)))) + (tclTRY (Proofview.V82.of_tactic (Equality.rewriteRL (mkVar eq_id)))) (do_rewrite eq_ids) g in @@ -412,31 +414,31 @@ let rewrite_until_var arg_num eq_ids : tactic = let rec_pte_id = Id.of_string "Hrec" let clean_hyp_with_heq ptes_infos eq_hyps hyp_id env sigma = - let coq_False = Coqlib.build_coq_False () in - let coq_True = Coqlib.build_coq_True () in - let coq_I = Coqlib.build_coq_I () in + let coq_False = EConstr.of_constr (Coqlib.build_coq_False ()) in + let coq_True = EConstr.of_constr (Coqlib.build_coq_True ()) in + let coq_I = EConstr.of_constr (Coqlib.build_coq_I ()) in let rec scan_type context type_of_hyp : tactic = - if isLetIn type_of_hyp then - let real_type_of_hyp = Term.it_mkProd_or_LetIn type_of_hyp context in + if isLetIn sigma type_of_hyp then + let real_type_of_hyp = it_mkProd_or_LetIn type_of_hyp context in let reduced_type_of_hyp = nf_betaiotazeta real_type_of_hyp in (* length of context didn't change ? *) let new_context,new_typ_of_hyp = - decompose_prod_n_assum (List.length context) reduced_type_of_hyp + decompose_prod_n_assum sigma (List.length context) reduced_type_of_hyp in tclTHENLIST [ h_reduce_with_zeta (Locusops.onHyp hyp_id); scan_type new_context new_typ_of_hyp ] - else if isProd type_of_hyp + else if isProd sigma type_of_hyp then begin - let (x,t_x,t') = destProd type_of_hyp in - let actual_real_type_of_hyp = Term.it_mkProd_or_LetIn t' context in - if is_property ptes_infos t_x actual_real_type_of_hyp then + let (x,t_x,t') = destProd sigma type_of_hyp in + let actual_real_type_of_hyp = it_mkProd_or_LetIn t' context in + if is_property sigma ptes_infos t_x actual_real_type_of_hyp then begin - let pte,pte_args = (destApp t_x) in - let (* fix_info *) prove_rec_hyp = (Id.Map.find (destVar pte) ptes_infos).proving_tac in + let pte,pte_args = (destApp sigma t_x) in + let (* fix_info *) prove_rec_hyp = (Id.Map.find (destVar sigma pte) ptes_infos).proving_tac in let popped_t' = pop t' in - let real_type_of_hyp = Term.it_mkProd_or_LetIn popped_t' context in + let real_type_of_hyp = it_mkProd_or_LetIn popped_t' context in let prove_new_type_of_hyp = let context_length = List.length context in tclTHENLIST @@ -453,8 +455,6 @@ let clean_hyp_with_heq ptes_infos eq_hyps hyp_id env sigma = List.rev_map mkVar (rec_pte_id::context_hyps_ids) ) in - let to_refine = EConstr.of_constr to_refine in - let t_x = EConstr.of_constr t_x in (* observe_tac "rec hyp " *) (tclTHENS (Proofview.V82.of_tactic (assert_before (Name rec_pte_id) t_x)) @@ -474,22 +474,22 @@ let clean_hyp_with_heq ptes_infos eq_hyps hyp_id env sigma = scan_type context popped_t' ] end - else if eq_constr t_x coq_False then + else if eq_constr sigma t_x coq_False then begin (* observe (str "Removing : "++ Ppconstr.pr_id hyp_id++ *) (* str " since it has False in its preconds " *) (* ); *) raise TOREMOVE; (* False -> .. useless *) end - else if is_incompatible_eq t_x then raise TOREMOVE (* t_x := C1 ... = C2 ... *) - else if eq_constr t_x coq_True (* Trivial => we remove this precons *) + else if is_incompatible_eq sigma t_x then raise TOREMOVE (* t_x := C1 ... = C2 ... *) + else if eq_constr sigma t_x coq_True (* Trivial => we remove this precons *) then (* observe (str "In "++Ppconstr.pr_id hyp_id++ *) (* str " removing useless precond True" *) (* ); *) let popped_t' = pop t' in let real_type_of_hyp = - Term.it_mkProd_or_LetIn popped_t' context + it_mkProd_or_LetIn popped_t' context in let prove_trivial = let nb_intro = List.length context in @@ -504,7 +504,6 @@ let clean_hyp_with_heq ptes_infos eq_hyps hyp_id env sigma = List.rev (coq_I::List.map mkVar context_hyps) ) in - let to_refine = (EConstr.of_constr to_refine) in refine to_refine g ) ] @@ -514,15 +513,15 @@ let clean_hyp_with_heq ptes_infos eq_hyps hyp_id env sigma = ((* observe_tac "prove_trivial" *) prove_trivial); scan_type context popped_t' ] - else if is_trivial_eq t_x + else if is_trivial_eq sigma t_x then (* t_x := t = t => we remove this precond *) let popped_t' = pop t' in let real_type_of_hyp = - Term.it_mkProd_or_LetIn popped_t' context + it_mkProd_or_LetIn popped_t' context in - let hd,args = destApp t_x in + let hd,args = destApp sigma t_x in let get_args hd args = - if eq_constr hd (Lazy.force eq) + if eq_constr sigma hd (Lazy.force eq) then (Lazy.force refl_equal,args.(0),args.(1)) else (jmeq_refl (),args.(0),args.(1)) in @@ -545,14 +544,14 @@ let clean_hyp_with_heq ptes_infos eq_hyps hyp_id env sigma = (scan_type new_context new_t') with Failure "NoChange" -> (* Last thing todo : push the rel in the context and continue *) - scan_type (LocalAssum (x,t_x) :: context) t' + scan_type (local_assum (x,t_x) :: context) t' end end else tclIDTAC in try - scan_type [] (Typing.unsafe_type_of env sigma (EConstr.mkVar hyp_id)), [hyp_id] + scan_type [] (EConstr.of_constr (Typing.unsafe_type_of env sigma (mkVar hyp_id))), [hyp_id] with TOREMOVE -> thin [hyp_id],[] @@ -602,26 +601,25 @@ let treat_new_case ptes_infos nb_prod continue_tac term dyn_infos = tclMAP (fun id -> Proofview.V82.of_tactic (introduction ~check:false id)) dyn_infos.rec_hyps; observe_tac "after_introduction" (fun g' -> (* We get infos on the equations introduced*) - let new_term_value_eq = pf_unsafe_type_of g' (EConstr.mkVar heq_id) in + let new_term_value_eq = pf_unsafe_type_of g' (mkVar heq_id) in + let new_term_value_eq = EConstr.of_constr new_term_value_eq in (* compute the new value of the body *) let new_term_value = - match kind_of_term new_term_value_eq with + match EConstr.kind (project g') new_term_value_eq with | App(f,[| _;_;args2 |]) -> args2 | _ -> observe (str "cannot compute new term value : " ++ pr_gls g' ++ fnl () ++ str "last hyp is" ++ - pr_lconstr_env (pf_env g') Evd.empty new_term_value_eq + pr_leconstr_env (pf_env g') (project g') new_term_value_eq ); anomaly (Pp.str "cannot compute new term value") in - let term = EConstr.of_constr term in let fun_body = mkLambda(Anonymous, - pf_unsafe_type_of g' term, - EConstr.Unsafe.to_constr (Termops.replace_term (project g') term (EConstr.mkRel 1) (EConstr.of_constr dyn_infos.info) - )) + EConstr.of_constr (pf_unsafe_type_of g' term), + Termops.replace_term (project g') term (mkRel 1) dyn_infos.info + ) in - let new_body = pf_nf_betaiota g' (EConstr.of_constr (mkApp(fun_body,[| new_term_value |]))) in - let new_body = EConstr.Unsafe.to_constr new_body in + let new_body = pf_nf_betaiota g' (mkApp(fun_body,[| new_term_value |])) in let new_infos = {dyn_infos with info = new_body; @@ -649,7 +647,6 @@ let instanciate_hyps_with_args (do_prove:Id.t list -> tactic) hyps args_id = fun g -> let prov_hid = pf_get_new_id hid g in let c = mkApp(mkVar hid,args) in - let c = EConstr.of_constr c in let evm, _ = pf_apply Typing.type_of g c in tclTHENLIST[ Refiner.tclEVARS evm; @@ -702,8 +699,9 @@ let build_proof : tactic = let rec build_proof_aux do_finalize dyn_infos : tactic = fun g -> + let sigma = project g in (* observe (str "proving on " ++ Printer.pr_lconstr_env (pf_env g) term);*) - match kind_of_term dyn_infos.info with + match EConstr.kind sigma dyn_infos.info with | Case(ci,ct,t,cb) -> let do_finalize_t dyn_info' = fun g -> @@ -711,18 +709,18 @@ let build_proof let dyn_infos = {dyn_info' with info = mkCase(ci,ct,t,cb)} in let g_nb_prod = nb_prod (project g) (EConstr.of_constr (pf_concl g)) in - let type_of_term = pf_unsafe_type_of g (EConstr.of_constr t) in + let type_of_term = pf_unsafe_type_of g t in + let type_of_term = EConstr.of_constr type_of_term in let term_eq = make_refl_eq (Lazy.force refl_equal) type_of_term t in - let term_eq = EConstr.of_constr term_eq in tclTHENSEQ [ - Proofview.V82.of_tactic (generalize (term_eq::(List.map EConstr.mkVar dyn_infos.rec_hyps))); + Proofview.V82.of_tactic (generalize (term_eq::(List.map mkVar dyn_infos.rec_hyps))); thin dyn_infos.rec_hyps; - Proofview.V82.of_tactic (pattern_option [Locus.AllOccurrencesBut [1],EConstr.of_constr t] None); + Proofview.V82.of_tactic (pattern_option [Locus.AllOccurrencesBut [1],t] None); (fun g -> observe_tac "toto" ( - tclTHENSEQ [Proofview.V82.of_tactic (Simple.case (EConstr.of_constr t)); + tclTHENSEQ [Proofview.V82.of_tactic (Simple.case t); (fun g' -> let g'_nb_prod = nb_prod (project g') (EConstr.of_constr (pf_concl g')) in let nb_instanciate_partial = g'_nb_prod - g_nb_prod in @@ -744,7 +742,7 @@ let build_proof build_proof do_finalize_t {dyn_infos with info = t} g | Lambda(n,t,b) -> begin - match kind_of_term( pf_concl g) with + match EConstr.kind sigma (EConstr.of_constr ( pf_concl g)) with | Prod _ -> tclTHEN (Proofview.V82.of_tactic intro) @@ -753,9 +751,8 @@ let build_proof let id = pf_last_hyp g' |> get_id in let new_term = pf_nf_betaiota g' - (EConstr.of_constr (mkApp(dyn_infos.info,[|mkVar id|]))) + (mkApp(dyn_infos.info,[|mkVar id|])) in - let new_term = EConstr.Unsafe.to_constr new_term in let new_infos = {dyn_infos with info = new_term} in let do_prove new_hyps = build_proof do_finalize @@ -775,9 +772,9 @@ let build_proof | Const _ | Var _ | Meta _ | Evar _ | Sort _ | Construct _ | Ind _ -> do_finalize dyn_infos g | App(_,_) -> - let f,args = decompose_app dyn_infos.info in + let f,args = decompose_app sigma dyn_infos.info in begin - match kind_of_term f with + match EConstr.kind sigma f with | App _ -> assert false (* we have collected all the app in decompose_app *) | Proj _ -> assert false (*FIXME*) | Var _ | Construct _ | Rel _ | Evar _ | Meta _ | Ind _ | Sort _ | Prod _ -> @@ -799,8 +796,7 @@ let build_proof do_finalize dyn_infos g | Lambda _ -> let new_term = - Reductionops.nf_beta Evd.empty (EConstr.of_constr dyn_infos.info) in - let new_term = EConstr.Unsafe.to_constr new_term in + Reductionops.nf_beta sigma dyn_infos.info in build_proof do_finalize {dyn_infos with info = new_term} g | LetIn _ -> @@ -852,7 +848,7 @@ let build_proof | Rel _ -> anomaly (Pp.str "Free var in goal conclusion !") and build_proof do_finalize dyn_infos g = (* observe (str "proving with "++Printer.pr_lconstr dyn_infos.info++ str " on goal " ++ pr_gls g); *) - observe_tac_stream (str "build_proof with " ++ Printer.pr_lconstr dyn_infos.info ) (build_proof_aux do_finalize dyn_infos) g + observe_tac_stream (str "build_proof with " ++ Printer.pr_leconstr dyn_infos.info ) (build_proof_aux do_finalize dyn_infos) g and build_proof_args do_finalize dyn_infos (* f_args' args *) :tactic = fun g -> let (f_args',args) = dyn_infos.info in @@ -918,11 +914,10 @@ let prove_rec_hyp_for_struct fix_info = (fun eq_hyps -> tclTHEN (rewrite_until_var (fix_info.idx) eq_hyps) (fun g -> - let _,pte_args = destApp (pf_concl g) in + let _,pte_args = destApp (project g) (EConstr.of_constr (pf_concl g)) in let rec_hyp_proof = mkApp(mkVar fix_info.name,array_get_start pte_args) in - let rec_hyp_proof = EConstr.of_constr rec_hyp_proof in refine rec_hyp_proof g )) @@ -936,7 +931,7 @@ let generalize_non_dep hyp g = (* observe (str "rec id := " ++ Ppconstr.pr_id hyp); *) let hyps = [hyp] in let env = Global.env () in - let hyp_typ = pf_unsafe_type_of g (EConstr.mkVar hyp) in + let hyp_typ = pf_unsafe_type_of g (mkVar hyp) in let to_revert,_ = let open Context.Named.Declaration in Environ.fold_named_context_reverse (fun (clear,keep) decl -> @@ -951,7 +946,7 @@ let generalize_non_dep hyp g = in (* observe (str "to_revert := " ++ prlist_with_sep spc Ppconstr.pr_id to_revert); *) tclTHEN - ((* observe_tac "h_generalize" *) (Proofview.V82.of_tactic (generalize (List.map EConstr.mkVar to_revert) ))) + ((* observe_tac "h_generalize" *) (Proofview.V82.of_tactic (generalize (List.map mkVar to_revert) ))) ((* observe_tac "thin" *) (thin to_revert)) g @@ -959,18 +954,19 @@ let id_of_decl = RelDecl.get_name %> Nameops.out_name let var_of_decl = id_of_decl %> mkVar let revert idl = tclTHEN - (Proofview.V82.of_tactic (generalize (List.map EConstr.mkVar idl))) + (Proofview.V82.of_tactic (generalize (List.map mkVar idl))) (thin idl) let generate_equation_lemma evd fnames f fun_num nb_params nb_args rec_args_num = (* observe (str "nb_args := " ++ str (string_of_int nb_args)); *) (* observe (str "nb_params := " ++ str (string_of_int nb_params)); *) (* observe (str "rec_args_num := " ++ str (string_of_int (rec_args_num + 1) )); *) - let f_def = Global.lookup_constant (fst (destConst f)) in + let f_def = Global.lookup_constant (fst (destConst evd f)) in let eq_lhs = mkApp(f,Array.init (nb_params + nb_args) (fun i -> mkRel(nb_params + nb_args - i))) in let f_body = Option.get (Global.body_of_constant_body f_def) in - let params,f_body_with_params = decompose_lam_n nb_params f_body in - let (_,num),(_,_,bodies) = destFix f_body_with_params in + let f_body = EConstr.of_constr f_body in + let params,f_body_with_params = decompose_lam_n evd nb_params f_body in + let (_,num),(_,_,bodies) = destFix evd f_body_with_params in let fnames_with_params = let params = Array.init nb_params (fun i -> mkRel(nb_params - i)) in let fnames = List.rev (Array.to_list (Array.map (fun f -> mkApp(f,params)) fnames)) in @@ -983,16 +979,15 @@ let generate_equation_lemma evd fnames f fun_num nb_params nb_args rec_args_num let eq_rhs = nf_betaiotazeta (mkApp(compose_lam params f_body_with_params_and_other_fun,Array.init (nb_params + nb_args) (fun i -> mkRel(nb_params + nb_args - i)))) in (* observe (str "eq_rhs " ++ pr_lconstr eq_rhs); *) let (type_ctxt,type_of_f),evd = - let evd,t = Typing.type_of ~refresh:true (Global.env ()) evd (EConstr.of_constr f) + let evd,t = Typing.type_of ~refresh:true (Global.env ()) evd f in - let t = EConstr.Unsafe.to_constr t in - decompose_prod_n_assum + decompose_prod_n_assum evd (nb_params + nb_args) t,evd in let eqn = mkApp(Lazy.force eq,[|type_of_f;eq_lhs;eq_rhs|]) in - let lemma_type = Term.it_mkProd_or_LetIn eqn type_ctxt in + let lemma_type = it_mkProd_or_LetIn eqn type_ctxt in (* Pp.msgnl (str "lemma type " ++ Printer.pr_lconstr lemma_type ++ fnl () ++ str "f_body " ++ Printer.pr_lconstr f_body); *) - let f_id = Label.to_id (con_label (fst (destConst f))) in + let f_id = Label.to_id (con_label (fst (destConst evd f))) in let prove_replacement = tclTHENSEQ [ @@ -1001,7 +996,7 @@ let generate_equation_lemma evd fnames f fun_num nb_params nb_args rec_args_num let rec_id = pf_nth_hyp_id g 1 in tclTHENSEQ [observe_tac "generalize_non_dep in generate_equation_lemma" (generalize_non_dep rec_id); - observe_tac "h_case" (Proofview.V82.of_tactic (simplest_case (EConstr.mkVar rec_id))); + observe_tac "h_case" (Proofview.V82.of_tactic (simplest_case (mkVar rec_id))); (Proofview.V82.of_tactic intros_reflexivity)] g ) ] @@ -1014,7 +1009,7 @@ let generate_equation_lemma evd fnames f fun_num nb_params nb_args rec_args_num (mk_equation_id f_id) (Decl_kinds.Global, Flags.is_universe_polymorphism (), (Decl_kinds.Proof Decl_kinds.Theorem)) evd - lemma_type + (EConstr.Unsafe.to_constr lemma_type) (Lemmas.mk_hook (fun _ _ -> ())); ignore (Pfedit.by (Proofview.V82.tactic prove_replacement)); Lemmas.save_proof (Vernacexpr.(Proved(Transparent,None))); @@ -1026,10 +1021,10 @@ let generate_equation_lemma evd fnames f fun_num nb_params nb_args rec_args_num let do_replace (evd:Evd.evar_map ref) params rec_arg_num rev_args_id f fun_num all_funs g = let equation_lemma = try - let finfos = find_Function_infos (fst (destConst f)) (*FIXME*) in + let finfos = find_Function_infos (fst (destConst !evd f)) (*FIXME*) in mkConst (Option.get finfos.equation_lemma) with (Not_found | Option.IsNone as e) -> - let f_id = Label.to_id (con_label (fst (destConst f))) in + let f_id = Label.to_id (con_label (fst (destConst !evd f))) in (*i The next call to mk_equation_id is valid since we will construct the lemma Ensures by: obvious i*) @@ -1038,7 +1033,7 @@ let do_replace (evd:Evd.evar_map ref) params rec_arg_num rev_args_id f fun_num a let _ = match e with | Option.IsNone -> - let finfos = find_Function_infos (fst (destConst f)) in + let finfos = find_Function_infos (fst (destConst !evd f)) in update_Function {finfos with equation_lemma = Some (match Nametab.locate (qualid_of_ident equation_lemma_id) with @@ -1054,8 +1049,9 @@ let do_replace (evd:Evd.evar_map ref) params rec_arg_num rev_args_id f fun_num a (Global.env ()) !evd (Constrintern.locate_reference (qualid_of_ident equation_lemma_id)) in + let res = EConstr.of_constr res in evd:=evd'; - let _ = Typing.e_type_of ~refresh:true (Global.env ()) evd (EConstr.of_constr res) in + let _ = Typing.e_type_of ~refresh:true (Global.env ()) evd res in res in let nb_intro_to_do = nb_prod (project g) (EConstr.of_constr (pf_concl g)) in @@ -1066,7 +1062,7 @@ let do_replace (evd:Evd.evar_map ref) params rec_arg_num rev_args_id f fun_num a let just_introduced = nLastDecls nb_intro_to_do g' in let open Context.Named.Declaration in let just_introduced_id = List.map get_id just_introduced in - tclTHEN (Proofview.V82.of_tactic (Equality.rewriteLR (EConstr.of_constr equation_lemma))) + tclTHEN (Proofview.V82.of_tactic (Equality.rewriteLR equation_lemma)) (revert just_introduced_id) g' ) g @@ -1103,15 +1099,15 @@ let prove_princ_for_struct (evd:Evd.evar_map ref) interactive_proof fun_num fnam let get_body const = match Global.body_of_constant const with | Some body -> - EConstr.Unsafe.to_constr (Tacred.cbv_norm_flags + Tacred.cbv_norm_flags (CClosure.RedFlags.mkflags [CClosure.RedFlags.fZETA]) (Global.env ()) (Evd.empty) - (EConstr.of_constr body)) + (EConstr.of_constr body) | None -> error ( "Cannot define a principle over an axiom ") in let fbody = get_body fnames.(fun_num) in - let f_ctxt,f_body = decompose_lam fbody in + let f_ctxt,f_body = decompose_lam (project g) fbody in let f_ctxt_length = List.length f_ctxt in let diff_params = princ_info.nparams - f_ctxt_length in let full_params,princ_params,fbody_with_full_params = @@ -1146,35 +1142,35 @@ let prove_princ_for_struct (evd:Evd.evar_map ref) interactive_proof fun_num fnam princ_params ); observe (str "fbody_with_full_params := " ++ - pr_lconstr fbody_with_full_params + pr_leconstr fbody_with_full_params ); let all_funs_with_full_params = Array.map (fun f -> applist(f, List.rev_map var_of_decl full_params)) all_funs in let fix_offset = List.length princ_params in let ptes_to_fix,infos = - match kind_of_term fbody_with_full_params with + match EConstr.kind (project g) fbody_with_full_params with | Fix((idxs,i),(names,typess,bodies)) -> let bodies_with_all_params = Array.map (fun body -> - EConstr.Unsafe.to_constr (Reductionops.nf_betaiota Evd.empty - (EConstr.of_constr (applist(substl (List.rev (Array.to_list all_funs_with_full_params)) body, - List.rev_map var_of_decl princ_params)))) + Reductionops.nf_betaiota (project g) + (applist(substl (List.rev (Array.to_list all_funs_with_full_params)) body, + List.rev_map var_of_decl princ_params)) ) bodies in let info_array = Array.mapi (fun i types -> - let types = Term.prod_applist types (List.rev_map var_of_decl princ_params) in + let types = prod_applist (project g) types (List.rev_map var_of_decl princ_params) in { idx = idxs.(i) - fix_offset; name = Nameops.out_name (fresh_id names.(i)); types = types; offset = fix_offset; nb_realargs = List.length - (fst (decompose_lam bodies.(i))) - fix_offset; + (fst (decompose_lam (project g) bodies.(i))) - fix_offset; body_with_param = bodies_with_all_params.(i); num_in_block = i } @@ -1186,7 +1182,7 @@ let prove_princ_for_struct (evd:Evd.evar_map ref) interactive_proof fun_num fnam (fun i (acc_map,acc_info) decl -> let pte = RelDecl.get_name decl in let infos = info_array.(i) in - let type_args,_ = decompose_prod infos.types in + let type_args,_ = decompose_prod (project g) infos.types in let nargs = List.length type_args in let f = applist(mkConst fnames.(i), List.rev_map var_of_decl princ_info.params) in let first_args = Array.init nargs (fun i -> mkRel (nargs -i)) in @@ -1196,20 +1192,20 @@ let prove_princ_for_struct (evd:Evd.evar_map ref) interactive_proof fun_num fnam let body_with_param,num = let body = get_body fnames.(i) in let body_with_full_params = - EConstr.Unsafe.to_constr (Reductionops.nf_betaiota Evd.empty (EConstr.of_constr ( - applist(body,List.rev_map var_of_decl full_params)))) + Reductionops.nf_betaiota (project g) ( + applist(body,List.rev_map var_of_decl full_params)) in - match kind_of_term body_with_full_params with + match EConstr.kind (project g) body_with_full_params with | Fix((_,num),(_,_,bs)) -> - EConstr.Unsafe.to_constr (Reductionops.nf_betaiota Evd.empty - (EConstr.of_constr ( + Reductionops.nf_betaiota (project g) + ( (applist (substl (List.rev (Array.to_list all_funs_with_full_params)) bs.(num), List.rev_map var_of_decl princ_params)) - ))),num + ),num | _ -> error "Not a mutual block" in let info = @@ -1238,7 +1234,7 @@ let prove_princ_for_struct (evd:Evd.evar_map ref) interactive_proof fun_num fnam | _, this_fix_info::others_infos -> let other_fix_infos = List.map - (fun fi -> fi.name,fi.idx + 1 ,EConstr.of_constr fi.types) + (fun fi -> fi.name,fi.idx + 1 ,fi.types) (pre_info@others_infos) in if List.is_empty other_fix_infos @@ -1262,11 +1258,11 @@ let prove_princ_for_struct (evd:Evd.evar_map ref) interactive_proof fun_num fnam in let intros_after_fixes : tactic = fun gl -> - let ctxt,pte_app = (decompose_prod_assum (pf_concl gl)) in - let pte,pte_args = (decompose_app pte_app) in + let ctxt,pte_app = (decompose_prod_assum (project gl) (EConstr.of_constr (pf_concl gl))) in + let pte,pte_args = (decompose_app (project gl) pte_app) in try let pte = - try destVar pte + try destVar (project gl) pte with DestKO -> anomaly (Pp.str "Property is not a variable") in let fix_info = Id.Map.find pte ptes_to_fix in @@ -1285,8 +1281,8 @@ let prove_princ_for_struct (evd:Evd.evar_map ref) interactive_proof fun_num fnam nb_rec_hyps = -100; rec_hyps = []; info = - EConstr.Unsafe.to_constr (Reductionops.nf_betaiota Evd.empty - (EConstr.of_constr (applist(fix_body,List.rev_map mkVar args_id)))); + Reductionops.nf_betaiota (project g) + (applist(fix_body,List.rev_map mkVar args_id)); eq_hyps = [] } in @@ -1345,15 +1341,15 @@ let prove_princ_for_struct (evd:Evd.evar_map ref) interactive_proof fun_num fnam nb_rec_hyps = -100; rec_hyps = []; info = - EConstr.Unsafe.to_constr (Reductionops.nf_betaiota Evd.empty - (EConstr.of_constr (applist(fbody_with_full_params, + Reductionops.nf_betaiota Evd.empty + (applist(fbody_with_full_params, (List.rev_map var_of_decl princ_params)@ (List.rev_map mkVar args_id) - )))); + )); eq_hyps = [] } in - let fname = destConst (fst (decompose_app (List.hd (List.rev pte_args)))) in + let fname = destConst (project g) (fst (decompose_app (project g) (List.hd (List.rev pte_args)))) in tclTHENSEQ [Proofview.V82.of_tactic (unfold_in_concl [(Locus.AllOccurrences, Names.EvalConstRef (fst fname))]); let do_prove = @@ -1431,18 +1427,18 @@ let prove_with_tcc tcc_lemma_constr eqs : tactic = let backtrack_eqs_until_hrec hrec eqs : tactic = fun gls -> - let eqs = List.map EConstr.mkVar eqs in + let eqs = List.map mkVar eqs in let rewrite = tclFIRST (List.map (fun x -> Proofview.V82.of_tactic (Equality.rewriteRL x)) eqs ) in - let _,hrec_concl = decompose_prod (pf_unsafe_type_of gls (EConstr.mkVar hrec)) in - let f_app = Array.last (snd (destApp hrec_concl)) in - let f = (fst (destApp f_app)) in + let _,hrec_concl = decompose_prod (project gls) (EConstr.of_constr (pf_unsafe_type_of gls (mkVar hrec))) in + let f_app = Array.last (snd (destApp (project gls) hrec_concl)) in + let f = (fst (destApp (project gls) f_app)) in let rec backtrack : tactic = fun g -> - let f_app = Array.last (snd (destApp (pf_concl g))) in - match kind_of_term f_app with - | App(f',_) when eq_constr f' f -> tclIDTAC g + let f_app = Array.last (snd (destApp (project g) (EConstr.of_constr (pf_concl g)))) in + match EConstr.kind (project g) f_app with + | App(f',_) when eq_constr (project g) f' f -> tclIDTAC g | _ -> tclTHEN rewrite backtrack g in backtrack gls @@ -1459,7 +1455,7 @@ let rec rewrite_eqs_in_eqs eqs = observe_tac (Format.sprintf "rewrite %s in %s " (Id.to_string eq) (Id.to_string id)) (tclTRY (Proofview.V82.of_tactic (Equality.general_rewrite_in true Locus.AllOccurrences - true (* dep proofs also: *) true id (EConstr.mkVar eq) false))) + true (* dep proofs also: *) true id (mkVar eq) false))) gl ) eqs @@ -1473,11 +1469,11 @@ let new_prove_with_tcc is_mes acc_inv hrec tcc_hyps eqs : tactic = backtrack_eqs_until_hrec hrec eqs; (* observe_tac ("new_prove_with_tcc ( applying "^(Id.to_string hrec)^" )" ) *) (tclTHENS (* We must have exactly ONE subgoal !*) - (Proofview.V82.of_tactic (apply (EConstr.mkVar hrec))) + (Proofview.V82.of_tactic (apply (mkVar hrec))) [ tclTHENSEQ [ (Proofview.V82.of_tactic (keep (tcc_hyps@eqs))); - (Proofview.V82.of_tactic (apply (EConstr.of_constr (Lazy.force acc_inv)))); + (Proofview.V82.of_tactic (apply (Lazy.force acc_inv))); (fun g -> if is_mes then @@ -1493,7 +1489,7 @@ let new_prove_with_tcc is_mes acc_inv hrec tcc_hyps eqs : tactic = tclCOMPLETE( Eauto.eauto_with_bases (true,5) - [{ Tacexpr.delayed = fun _ sigma -> Sigma.here (EConstr.of_constr (Lazy.force refl_equal)) sigma}] + [{ Tacexpr.delayed = fun _ sigma -> Sigma.here (Lazy.force refl_equal) sigma}] [Hints.Hint_db.empty empty_transparent_state false] ) ) @@ -1506,20 +1502,20 @@ let new_prove_with_tcc is_mes acc_inv hrec tcc_hyps eqs : tactic = gls -let is_valid_hypothesis predicates_name = +let is_valid_hypothesis sigma predicates_name = let predicates_name = List.fold_right Id.Set.add predicates_name Id.Set.empty in let is_pte typ = - if isApp typ + if isApp sigma typ then - let pte,_ = destApp typ in - if isVar pte - then Id.Set.mem (destVar pte) predicates_name + let pte,_ = destApp sigma typ in + if isVar sigma pte + then Id.Set.mem (destVar sigma pte) predicates_name else false else false in let rec is_valid_hypothesis typ = is_pte typ || - match kind_of_term typ with + match EConstr.kind sigma typ with | Prod(_,pte,typ') -> is_pte pte && is_valid_hypothesis typ' | _ -> false in @@ -1584,7 +1580,7 @@ let prove_principle_for_gen Nameops.out_name (fresh_id (Name (Id.of_string ("Acc_"^(Id.to_string rec_arg_id))))) in let revert l = - tclTHEN (Proofview.V82.of_tactic (Tactics.generalize (List.map EConstr.mkVar l))) (Proofview.V82.of_tactic (clear l)) + tclTHEN (Proofview.V82.of_tactic (Tactics.generalize (List.map mkVar l))) (Proofview.V82.of_tactic (clear l)) in let fix_id = Nameops.out_name (fresh_id (Name hrec_id)) in let prove_rec_arg_acc g = @@ -1592,12 +1588,12 @@ let prove_principle_for_gen (tclCOMPLETE (tclTHEN (Proofview.V82.of_tactic (assert_by (Name wf_thm_id) - (EConstr.of_constr (mkApp (delayed_force well_founded,[|input_type;relation|]))) + (mkApp (delayed_force well_founded,[|input_type;relation|])) (Proofview.V82.tactic (fun g -> (* observe_tac "prove wf" *) (tclCOMPLETE (wf_tac is_mes)) g)))) ( (* observe_tac *) (* "apply wf_thm" *) - Proofview.V82.of_tactic (Tactics.Simple.apply (EConstr.of_constr (mkApp(mkVar wf_thm_id,[|mkVar rec_arg_id|])))) + Proofview.V82.of_tactic (Tactics.Simple.apply (mkApp(mkVar wf_thm_id,[|mkVar rec_arg_id|]))) ) ) ) @@ -1632,7 +1628,7 @@ let prove_principle_for_gen [ Proofview.V82.of_tactic (generalize [lemma]); Proofview.V82.of_tactic (Simple.intro hid); - Proofview.V82.of_tactic (Elim.h_decompose_and (EConstr.mkVar hid)); + Proofview.V82.of_tactic (Elim.h_decompose_and (mkVar hid)); (fun g -> let new_hyps = pf_ids_of_hyps g in tcc_list := List.rev (List.subtract Id.equal new_hyps (hid::hyps)); @@ -1656,7 +1652,7 @@ let prove_principle_for_gen ); (* observe_tac "" *) Proofview.V82.of_tactic (assert_by (Name acc_rec_arg_id) - (EConstr.of_constr (mkApp (delayed_force acc_rel,[|input_type;relation;mkVar rec_arg_id|]))) + (mkApp (delayed_force acc_rel,[|input_type;relation;mkVar rec_arg_id|])) (Proofview.V82.tactic prove_rec_arg_acc) ); (* observe_tac "reverting" *) (revert (List.rev (acc_rec_arg_id::args_ids))); @@ -1665,10 +1661,10 @@ let prove_principle_for_gen (* observe_tac "h_fix " *) (Proofview.V82.of_tactic (fix (Some fix_id) (List.length args_ids + 1))); (* (fun g -> observe (Printer.pr_goal (sig_it g) ++ fnl() ++ pr_lconstr_env (pf_env g ) (pf_unsafe_type_of g (mkVar fix_id) )); tclIDTAC g); *) h_intros (List.rev (acc_rec_arg_id::args_ids)); - Proofview.V82.of_tactic (Equality.rewriteLR (EConstr.of_constr (mkConst eq_ref))); + Proofview.V82.of_tactic (Equality.rewriteLR (mkConst eq_ref)); (* observe_tac "finish" *) (fun gl' -> let body = - let _,args = destApp (pf_concl gl') in + let _,args = destApp (project gl') (EConstr.of_constr (pf_concl gl')) in Array.last args in let body_info rec_hyps = @@ -1711,7 +1707,7 @@ let prove_principle_for_gen ) ); - is_valid = is_valid_hypothesis predicates_names + is_valid = is_valid_hypothesis (project gl') predicates_names } in let ptes_info : pte_info Id.Map.t = diff --git a/plugins/funind/functional_principles_proofs.mli b/plugins/funind/functional_principles_proofs.mli index 34ce66967..769d726d7 100644 --- a/plugins/funind/functional_principles_proofs.mli +++ b/plugins/funind/functional_principles_proofs.mli @@ -4,7 +4,7 @@ open Term val prove_princ_for_struct : Evd.evar_map ref -> bool -> - int -> constant array -> constr array -> int -> Tacmach.tactic + int -> constant array -> EConstr.constr array -> int -> Tacmach.tactic val prove_principle_for_gen : @@ -12,8 +12,8 @@ val prove_principle_for_gen : constr option ref -> (* a pointer to the obligation proofs lemma *) bool -> (* is that function uses measure *) int -> (* the number of recursive argument *) - types -> (* the type of the recursive argument *) - constr -> (* the wf relation used to prove the function *) + EConstr.types -> (* the type of the recursive argument *) + EConstr.constr -> (* the wf relation used to prove the function *) Tacmach.tactic diff --git a/plugins/funind/functional_principles_types.ml b/plugins/funind/functional_principles_types.ml index 8683f68c6..d964002f9 100644 --- a/plugins/funind/functional_principles_types.ml +++ b/plugins/funind/functional_principles_types.ml @@ -63,7 +63,7 @@ let compute_new_princ_type_from_rel rel_to_fun sorts princ_type = else args in Context.Named.Declaration.LocalAssum (Nameops.out_name (Context.Rel.Declaration.get_name decl), - compose_prod real_args (mkSort new_sort)) + Term.compose_prod real_args (mkSort new_sort)) in let new_predicates = List.map_i @@ -254,7 +254,7 @@ let change_property_sort evd toSort princ princName = let args,ty = decompose_prod (get_type decl) in let s = destSort ty in Global.add_constraints (Univ.enforce_leq (univ_of_sort toSort) (univ_of_sort s) Univ.Constraint.empty); - compose_prod args (mkSort toSort) + Term.compose_prod args (mkSort toSort) ) in let evd,princName_as_constr = @@ -298,7 +298,7 @@ let build_functional_principle (evd:Evd.evar_map ref) interactive_proof old_prin hook ; (* let _tim1 = System.get_time () in *) - ignore (Pfedit.by (Proofview.V82.tactic (proof_tac (Array.map mkConstU funs) mutr_nparams))); + ignore (Pfedit.by (Proofview.V82.tactic (proof_tac (Array.map EConstr.mkConstU funs) mutr_nparams))); (* let _tim2 = System.get_time () in *) (* begin *) (* let dur1 = System.time_difference tim1 tim2 in *) diff --git a/plugins/funind/functional_principles_types.mli b/plugins/funind/functional_principles_types.mli index 3fa2644ca..45ad332fc 100644 --- a/plugins/funind/functional_principles_types.mli +++ b/plugins/funind/functional_principles_types.mli @@ -27,7 +27,7 @@ val generate_functional_principle : (* The tactic to use to make the proof w.r the number of params *) - (constr array -> int -> Tacmach.tactic) -> + (EConstr.constr array -> int -> Tacmach.tactic) -> unit val compute_new_princ_type_from_rel : constr array -> sorts array -> diff --git a/plugins/funind/g_indfun.ml4 b/plugins/funind/g_indfun.ml4 index 560242bf2..27a892ca7 100644 --- a/plugins/funind/g_indfun.ml4 +++ b/plugins/funind/g_indfun.ml4 @@ -98,7 +98,6 @@ ARGUMENT EXTEND with_names TYPED AS intropattern_opt PRINTED BY pr_intro_as_pat END let functional_induction b c x pat = - let x = Option.map (Miscops.map_with_bindings EConstr.Unsafe.to_constr) x in Proofview.V82.tactic (functional_induction true c x (Option.map out_disjunctive pat)) @@ -110,7 +109,6 @@ TACTIC EXTEND newfunind | [c] -> c | c::cl -> EConstr.applist(c,cl) in - let c = EConstr.Unsafe.to_constr c in Extratactics.onSomeWithHoles (fun x -> functional_induction true c x pat) princl ] END (***** debug only ***) @@ -122,7 +120,6 @@ TACTIC EXTEND snewfunind | [c] -> c | c::cl -> EConstr.applist(c,cl) in - let c = EConstr.Unsafe.to_constr c in Extratactics.onSomeWithHoles (fun x -> functional_induction false c x pat) princl ] END diff --git a/plugins/funind/glob_term_to_relation.ml b/plugins/funind/glob_term_to_relation.ml index 0725bb11c..fc5a287ae 100644 --- a/plugins/funind/glob_term_to_relation.ml +++ b/plugins/funind/glob_term_to_relation.ml @@ -952,7 +952,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt = in mkGProd(n,t,new_b),id_to_exclude with Continue -> - let jmeq = Globnames.IndRef (fst (destInd (jmeq ()))) in + let jmeq = Globnames.IndRef (fst (EConstr.destInd Evd.empty (jmeq ()))) in let ty',ctx = Pretyping.understand env (Evd.from_env env) ty in let ind,args' = Inductive.find_inductive env ty' in let mib,_ = Global.lookup_inductive (fst ind) in diff --git a/plugins/funind/indfun.ml b/plugins/funind/indfun.ml index 1b899c152..e22fed391 100644 --- a/plugins/funind/indfun.ml +++ b/plugins/funind/indfun.ml @@ -2,6 +2,7 @@ open CErrors open Util open Names open Term +open EConstr open Pp open Indfun_common open Libnames @@ -18,8 +19,8 @@ let is_rec_info sigma scheme_info = let test_branche min acc decl = acc || ( let new_branche = - it_mkProd_or_LetIn mkProp (fst (decompose_prod_assum (RelDecl.get_type decl))) in - let free_rels_in_br = Termops.free_rels sigma (EConstr.of_constr new_branche) in + it_mkProd_or_LetIn mkProp (fst (decompose_prod_assum sigma (EConstr.of_constr (RelDecl.get_type decl)))) in + let free_rels_in_br = Termops.free_rels sigma new_branche in let max = min + scheme_info.Tactics.npredicates in Int.Set.exists (fun i -> i >= min && i< max) free_rels_in_br ) @@ -32,20 +33,21 @@ let choose_dest_or_ind scheme_info args = let functional_induction with_clean c princl pat = let res = - let f,args = decompose_app c in fun g -> + let sigma = Tacmach.project g in + let f,args = decompose_app sigma c in let princ,bindings, princ_type,g' = match princl with | None -> (* No principle is given let's find the good one *) begin - match kind_of_term f with + match EConstr.kind sigma f with | Const (c',u) -> let princ_option = let finfo = (* we first try to find out a graph on f *) try find_Function_infos c' with Not_found -> user_err (str "Cannot find induction information on "++ - Printer.pr_lconstr (mkConst c') ) + Printer.pr_leconstr (mkConst c') ) in match Tacticals.elimination_sort_of_goal g with | InProp -> finfo.prop_lemma @@ -73,15 +75,16 @@ let functional_induction with_clean c princl pat = (* mkConst(const_of_id princ_name ),g (\* FIXME *\) *) with Not_found -> (* This one is neither defined ! *) user_err (str "Cannot find induction principle for " - ++Printer.pr_lconstr (mkConst c') ) + ++Printer.pr_leconstr (mkConst c') ) in - (princ,NoBindings, Tacmach.pf_unsafe_type_of g' (EConstr.of_constr princ),g') + let princ = EConstr.of_constr princ in + (princ,NoBindings,EConstr.of_constr (Tacmach.pf_unsafe_type_of g' princ),g') | _ -> raise (UserError(None,str "functional induction must be used with a function" )) end | Some ((princ,binding)) -> - princ,binding,Tacmach.pf_unsafe_type_of g (EConstr.of_constr princ),g + princ,binding,EConstr.of_constr (Tacmach.pf_unsafe_type_of g princ),g in - let princ_type = EConstr.of_constr princ_type in + let sigma = Tacmach.project g' in let princ_infos = Tactics.compute_elim_sig (Tacmach.project g') princ_type in let args_as_induction_constr = let c_list = @@ -90,15 +93,13 @@ let functional_induction with_clean c princl pat = in let encoded_pat_as_patlist = List.make (List.length args + List.length c_list - 1) None @ [pat] in - List.map2 (fun c pat -> ((None,Tacexpr.ElimOnConstr ({ Tacexpr.delayed = fun env sigma -> Sigma ((EConstr.of_constr c,NoBindings), sigma, Sigma.refl) })),(None,pat),None)) + List.map2 (fun c pat -> ((None,Tacexpr.ElimOnConstr ({ Tacexpr.delayed = fun env sigma -> Sigma ((c,NoBindings), sigma, Sigma.refl) })),(None,pat),None)) (args@c_list) encoded_pat_as_patlist in - let princ = EConstr.of_constr princ in - let bindings = Miscops.map_bindings EConstr.of_constr bindings in let princ' = Some (princ,bindings) in let princ_vars = List.fold_right - (fun a acc -> try Id.Set.add (destVar a) acc with DestKO -> acc) + (fun a acc -> try Id.Set.add (destVar sigma a) acc with DestKO -> acc) args Id.Set.empty in @@ -247,7 +248,8 @@ let derive_inversion fix_names = let evd,c = Evd.fresh_global (Global.env ()) evd (Constrintern.locate_reference (Libnames.qualid_of_ident id)) in - evd, destConst c::l + let c = EConstr.of_constr c in + evd, destConst evd c::l ) fix_names (evd',[]) @@ -267,7 +269,8 @@ let derive_inversion fix_names = (Global.env ()) evd (Constrintern.locate_reference (Libnames.qualid_of_ident (mk_rel_id id))) in - evd,(fst (destInd id))::l + let id = EConstr.of_constr id in + evd,(fst (destInd evd id))::l ) fix_names (evd',[]) @@ -334,7 +337,7 @@ let error_error names e = let generate_principle (evd:Evd.evar_map ref) pconstants on_error is_general do_built (fix_rec_l:(Vernacexpr.fixpoint_expr * Vernacexpr.decl_notation list) list) recdefs interactive_proof - (continue_proof : int -> Names.constant array -> Term.constr array -> int -> + (continue_proof : int -> Names.constant array -> EConstr.constr array -> int -> Tacmach.tactic) : unit = let names = List.map (function (((_, name),_),_,_,_,_),_ -> name) fix_rec_l in let fun_bodies = List.map2 prepare_body fix_rec_l recdefs in @@ -408,7 +411,8 @@ let register_struct is_rec (fixpoint_exprl:(Vernacexpr.fixpoint_expr * Vernacexp let evd,c = Evd.fresh_global (Global.env ()) evd (Constrintern.locate_reference (Libnames.qualid_of_ident fname)) in - evd,((destConst c)::l) + let c = EConstr.of_constr c in + evd,((destConst evd c)::l) ) (Evd.from_env (Global.env ()),[]) fixpoint_exprl @@ -422,7 +426,8 @@ let register_struct is_rec (fixpoint_exprl:(Vernacexpr.fixpoint_expr * Vernacexp let evd,c = Evd.fresh_global (Global.env ()) evd (Constrintern.locate_reference (Libnames.qualid_of_ident fname)) in - evd,((destConst c)::l) + let c = EConstr.of_constr c in + evd,((destConst evd c)::l) ) (Evd.from_env (Global.env ()),[]) fixpoint_exprl @@ -432,7 +437,7 @@ let register_struct is_rec (fixpoint_exprl:(Vernacexpr.fixpoint_expr * Vernacexp let generate_correction_proof_wf f_ref tcc_lemma_ref is_mes functional_ref eq_ref rec_arg_num rec_arg_type nb_args relation - (_: int) (_:Names.constant array) (_:Term.constr array) (_:int) : Tacmach.tactic = + (_: int) (_:Names.constant array) (_:EConstr.constr array) (_:int) : Tacmach.tactic = Functional_principles_proofs.prove_principle_for_gen (f_ref,functional_ref,eq_ref) tcc_lemma_ref is_mes rec_arg_num rec_arg_type relation @@ -840,7 +845,7 @@ let make_graph (f_ref:global_reference) = | ConstRef c -> begin try c,Global.lookup_constant c with Not_found -> - raise (UserError (None,str "Cannot find " ++ Printer.pr_lconstr (mkConst c)) ) + raise (UserError (None,str "Cannot find " ++ Printer.pr_leconstr (mkConst c)) ) end | _ -> raise (UserError (None, str "Not a function reference") ) in diff --git a/plugins/funind/indfun.mli b/plugins/funind/indfun.mli index 1c27bdfac..ba89fe4a7 100644 --- a/plugins/funind/indfun.mli +++ b/plugins/funind/indfun.mli @@ -12,8 +12,8 @@ val do_generate_principle : val functional_induction : bool -> - Term.constr -> - (Term.constr * Term.constr bindings) option -> + EConstr.constr -> + (EConstr.constr * EConstr.constr bindings) option -> Tacexpr.or_and_intro_pattern option -> Proof_type.goal Tacmach.sigma -> Proof_type.goal list Evd.sigma diff --git a/plugins/funind/indfun_common.ml b/plugins/funind/indfun_common.ml index 08b40a1f7..2889d8d03 100644 --- a/plugins/funind/indfun_common.ml +++ b/plugins/funind/indfun_common.ml @@ -130,8 +130,8 @@ let find_reference sl s = let dp = Names.DirPath.make (List.rev_map Id.of_string sl) in Nametab.locate (make_qualid dp (Id.of_string s)) -let eq = lazy(coq_constant "eq") -let refl_equal = lazy(coq_constant "eq_refl") +let eq = lazy(EConstr.of_constr (coq_constant "eq")) +let refl_equal = lazy(EConstr.of_constr (coq_constant "eq_refl")) (*****************************************************************) (* Copy of the standart save mechanism but without the much too *) @@ -475,13 +475,13 @@ exception ToShow of exn let jmeq () = try Coqlib.check_required_library Coqlib.jmeq_module_name; - Coqlib.gen_constant "Function" ["Logic";"JMeq"] "JMeq" + EConstr.of_constr (Coqlib.gen_constant "Function" ["Logic";"JMeq"] "JMeq") with e when CErrors.noncritical e -> raise (ToShow e) let jmeq_refl () = try Coqlib.check_required_library Coqlib.jmeq_module_name; - Coqlib.gen_constant "Function" ["Logic";"JMeq"] "JMeq_refl" + EConstr.of_constr (Coqlib.gen_constant "Function" ["Logic";"JMeq"] "JMeq_refl") with e when CErrors.noncritical e -> raise (ToShow e) let h_intros l = @@ -489,10 +489,10 @@ let h_intros l = let h_id = Id.of_string "h" let hrec_id = Id.of_string "hrec" -let well_founded = function () -> (coq_constant "well_founded") -let acc_rel = function () -> (coq_constant "Acc") -let acc_inv_id = function () -> (coq_constant "Acc_inv") -let well_founded_ltof = function () -> (Coqlib.coq_constant "" ["Arith";"Wf_nat"] "well_founded_ltof") +let well_founded = function () -> EConstr.of_constr (coq_constant "well_founded") +let acc_rel = function () -> EConstr.of_constr (coq_constant "Acc") +let acc_inv_id = function () -> EConstr.of_constr (coq_constant "Acc_inv") +let well_founded_ltof = function () -> EConstr.of_constr (Coqlib.coq_constant "" ["Arith";"Wf_nat"] "well_founded_ltof") let ltof_ref = function () -> (find_reference ["Coq";"Arith";"Wf_nat"] "ltof") let evaluable_of_global_reference r = (* Tacred.evaluable_of_global_reference (Global.env ()) *) @@ -501,9 +501,45 @@ let evaluable_of_global_reference r = (* Tacred.evaluable_of_global_reference (G | VarRef id -> EvalVarRef id | _ -> assert false;; -let list_rewrite (rev:bool) (eqs: (constr*bool) list) = - let eqs = List.map (Util.on_fst EConstr.of_constr) eqs in +let list_rewrite (rev:bool) (eqs: (EConstr.constr*bool) list) = tclREPEAT (List.fold_right (fun (eq,b) i -> tclORELSE (Proofview.V82.of_tactic ((if b then Equality.rewriteLR else Equality.rewriteRL) eq)) i) (if rev then (List.rev eqs) else eqs) (tclFAIL 0 (mt())));; + +let decompose_lam_n sigma n = + let open EConstr in + if n < 0 then CErrors.error "decompose_lam_n: integer parameter must be positive"; + let rec lamdec_rec l n c = + if Int.equal n 0 then l,c + else match EConstr.kind sigma c with + | Lambda (x,t,c) -> lamdec_rec ((x,t)::l) (n-1) c + | Cast (c,_,_) -> lamdec_rec l n c + | _ -> CErrors.error "decompose_lam_n: not enough abstractions" + in + lamdec_rec [] n + +let lamn n env b = + let open EConstr in + 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) + +(* compose_lam [xn:Tn;..;x1:T1] b = [x1:T1]..[xn:Tn]b *) +let compose_lam l b = lamn (List.length l) l b + +(* prodn n [xn:Tn;..;x1:T1;Gamma] b = (x1:T1)..(xn:Tn)b *) +let prodn n env b = + let open EConstr in + let rec prodrec = function + | (0, env, b) -> b + | (n, ((v,t)::l), b) -> prodrec (n-1, l, mkProd (v,t,b)) + | _ -> assert false + in + prodrec (n,env,b) + +(* compose_prod [xn:Tn;..;x1:T1] b = (x1:T1)..(xn:Tn)b *) +let compose_prod l b = prodn (List.length l) l b diff --git a/plugins/funind/indfun_common.mli b/plugins/funind/indfun_common.mli index e5c756f56..5836d6519 100644 --- a/plugins/funind/indfun_common.mli +++ b/plugins/funind/indfun_common.mli @@ -40,11 +40,11 @@ val chop_rprod_n : int -> Glob_term.glob_constr -> (Name.t*Glob_term.glob_constr) list * Glob_term.glob_constr val def_of_const : Term.constr -> Term.constr -val eq : Term.constr Lazy.t -val refl_equal : Term.constr Lazy.t +val eq : EConstr.constr Lazy.t +val refl_equal : EConstr.constr Lazy.t val const_of_id: Id.t -> Globnames.global_reference(* constantyes *) -val jmeq : unit -> Term.constr -val jmeq_refl : unit -> Term.constr +val jmeq : unit -> EConstr.constr +val jmeq_refl : unit -> EConstr.constr val save : bool -> Id.t -> Safe_typing.private_constants Entries.definition_entry -> Decl_kinds.goal_kind -> unit Lemmas.declaration_hook CEphemeron.key -> unit @@ -107,10 +107,15 @@ val is_strict_tcc : unit -> bool val h_intros: Names.Id.t list -> Proof_type.tactic val h_id : Names.Id.t val hrec_id : Names.Id.t -val acc_inv_id : Term.constr Util.delayed +val acc_inv_id : EConstr.constr Util.delayed val ltof_ref : Globnames.global_reference Util.delayed -val well_founded_ltof : Term.constr Util.delayed -val acc_rel : Term.constr Util.delayed -val well_founded : Term.constr Util.delayed +val well_founded_ltof : EConstr.constr Util.delayed +val acc_rel : EConstr.constr Util.delayed +val well_founded : EConstr.constr Util.delayed val evaluable_of_global_reference : Globnames.global_reference -> Names.evaluable_global_reference -val list_rewrite : bool -> (Term.constr*bool) list -> Proof_type.tactic +val list_rewrite : bool -> (EConstr.constr*bool) list -> Proof_type.tactic + +val decompose_lam_n : Evd.evar_map -> int -> EConstr.t -> + (Names.Name.t * EConstr.t) list * EConstr.t +val compose_lam : (Names.Name.t * EConstr.t) list -> EConstr.t -> EConstr.t +val compose_prod : (Names.Name.t * EConstr.t) list -> EConstr.t -> EConstr.t diff --git a/plugins/funind/invfun.ml b/plugins/funind/invfun.ml index 27528c2dc..be82010d9 100644 --- a/plugins/funind/invfun.ml +++ b/plugins/funind/invfun.ml @@ -12,6 +12,7 @@ open CErrors open Util open Names open Term +open EConstr open Vars open Pp open Globnames @@ -25,6 +26,12 @@ open Context.Rel.Declaration module RelDecl = Context.Rel.Declaration +let local_assum (na, t) = + RelDecl.LocalAssum (na, EConstr.Unsafe.to_constr t) + +let local_def (na, b, t) = + RelDecl.LocalDef (na, EConstr.Unsafe.to_constr b, EConstr.Unsafe.to_constr t) + (* Some pretty printing function for debugging purpose *) let pr_binding prc = @@ -108,11 +115,11 @@ let thin ids gl = Proofview.V82.of_tactic (Tactics.clear ids) gl let make_eq () = try - Universes.constr_of_global (Coqlib.build_coq_eq ()) + EConstr.of_constr (Universes.constr_of_global (Coqlib.build_coq_eq ())) with _ -> assert false let make_eq_refl () = try - Universes.constr_of_global (Coqlib.build_coq_eq_refl ()) + EConstr.of_constr (Universes.constr_of_global (Coqlib.build_coq_eq_refl ())) with _ -> assert false @@ -131,16 +138,16 @@ let make_eq_refl () = let generate_type evd g_to_f f graph i = (*i we deduce the number of arguments of the function and its returned type from the graph i*) let evd',graph = - Evd.fresh_global (Global.env ()) !evd (Globnames.IndRef (fst (destInd graph))) + Evd.fresh_global (Global.env ()) !evd (Globnames.IndRef (fst (destInd !evd graph))) in + let graph = EConstr.of_constr graph in evd:=evd'; - let graph_arity = Typing.e_type_of (Global.env ()) evd (EConstr.of_constr graph) in - let graph_arity = EConstr.Unsafe.to_constr graph_arity in - let ctxt,_ = decompose_prod_assum graph_arity in + let graph_arity = Typing.e_type_of (Global.env ()) evd graph in + let ctxt,_ = decompose_prod_assum !evd graph_arity in let fun_ctxt,res_type = match ctxt with | [] | [_] -> anomaly (Pp.str "Not a valid context") - | decl :: fun_ctxt -> fun_ctxt, RelDecl.get_type decl + | decl :: fun_ctxt -> fun_ctxt, EConstr.of_constr (RelDecl.get_type decl) in let rec args_from_decl i accu = function | [] -> accu @@ -180,12 +187,12 @@ let generate_type evd g_to_f f graph i = \[\forall (x_1:t_1)\ldots(x_n:t_n), let fv := f x_1\ldots x_n in, forall res, \] i*) let pre_ctxt = - LocalAssum (Name res_id, lift 1 res_type) :: LocalDef (Name fv_id, mkApp (f,args_as_rels), res_type) :: fun_ctxt + local_assum (Name res_id, lift 1 res_type) :: local_def (Name fv_id, mkApp (f,args_as_rels), res_type) :: fun_ctxt in (*i and we can return the solution depending on which lemma type we are defining i*) if g_to_f - then LocalAssum (Anonymous,graph_applied)::pre_ctxt,(lift 1 res_eq_f_of_args),graph - else LocalAssum (Anonymous,res_eq_f_of_args)::pre_ctxt,(lift 1 graph_applied),graph + then local_assum (Anonymous,graph_applied)::pre_ctxt,(lift 1 res_eq_f_of_args),graph + else local_assum (Anonymous,res_eq_f_of_args)::pre_ctxt,(lift 1 graph_applied),graph (* @@ -194,7 +201,7 @@ let generate_type evd g_to_f f graph i = WARNING: while convertible, [type_of body] and [type] can be non equal *) let find_induction_principle evd f = - let f_as_constant,u = match kind_of_term f with + let f_as_constant,u = match EConstr.kind !evd f with | Const c' -> c' | _ -> error "Must be used with a function" in @@ -203,8 +210,8 @@ let find_induction_principle evd f = | None -> raise Not_found | Some rect_lemma -> let evd',rect_lemma = Evd.fresh_global (Global.env ()) !evd (Globnames.ConstRef rect_lemma) in - let evd',typ = Typing.type_of ~refresh:true (Global.env ()) evd' (EConstr.of_constr rect_lemma) in - let typ = EConstr.Unsafe.to_constr typ in + let rect_lemma = EConstr.of_constr rect_lemma in + let evd',typ = Typing.type_of ~refresh:true (Global.env ()) evd' rect_lemma in evd:=evd'; rect_lemma,typ @@ -248,12 +255,12 @@ let prove_fun_correct evd functional_induction funs_constr graphs_constr schemes \[fun (x_1:t_1)\ldots(x_n:t_n)=> fun fv => fun res => res = fv \rightarrow graph\ x_1\ldots x_n\ res\] *) (* we the get the definition of the graphs block *) - let graph_ind,u = destInd graphs_constr.(i) in + let graph_ind,u = destInd evd graphs_constr.(i) in let kn = fst graph_ind in let mib,_ = Global.lookup_inductive graph_ind in (* and the principle to use in this lemma in $\zeta$ normal form *) let f_principle,princ_type = schemes.(i) in - let princ_type = nf_zeta (EConstr.of_constr princ_type) in + let princ_type = nf_zeta princ_type in let princ_infos = Tactics.compute_elim_sig evd princ_type in (* The number of args of the function is then easily computable *) let nb_fun_args = nb_prod (project g) (EConstr.of_constr (pf_concl g)) - 2 in @@ -273,13 +280,13 @@ let prove_fun_correct evd functional_induction funs_constr graphs_constr schemes (fun decl -> List.map (fun id -> Loc.ghost, IntroNaming (IntroIdentifier id)) - (generate_fresh_id (Id.of_string "y") ids (List.length (fst (decompose_prod_assum (RelDecl.get_type decl))))) + (generate_fresh_id (Id.of_string "y") ids (List.length (fst (decompose_prod_assum evd (EConstr.of_constr (RelDecl.get_type decl)))))) ) branches in (* before building the full intro pattern for the principle *) let eq_ind = make_eq () in - let eq_construct = mkConstructUi (destInd eq_ind, 1) in + let eq_construct = mkConstructUi (destInd evd eq_ind, 1) in (* The next to referencies will be used to find out which constructor to apply in each branch *) let ind_number = ref 0 and min_constr_number = ref 0 in @@ -307,18 +314,20 @@ let prove_fun_correct evd functional_induction funs_constr graphs_constr schemes let constructor_args g = List.fold_right (fun hid acc -> - let type_of_hid = pf_unsafe_type_of g (EConstr.mkVar hid) in - match kind_of_term type_of_hid with + let type_of_hid = pf_unsafe_type_of g (mkVar hid) in + let type_of_hid = EConstr.of_constr type_of_hid in + let sigma = project g in + match EConstr.kind sigma type_of_hid with | Prod(_,_,t') -> begin - match kind_of_term t' with + match EConstr.kind sigma t' with | Prod(_,t'',t''') -> begin - match kind_of_term t'',kind_of_term t''' with + match EConstr.kind sigma t'',EConstr.kind sigma t''' with | App(eq,args), App(graph',_) when - (Term.eq_constr eq eq_ind) && - Array.exists (Constr.eq_constr_nounivs graph') graphs_constr -> + (EConstr.eq_constr sigma eq eq_ind) && + Array.exists (EConstr.eq_constr_nounivs sigma graph') graphs_constr -> (args.(2)::(mkApp(mkVar hid,[|args.(2);(mkApp(eq_construct,[|args.(0);args.(2)|]))|])) ::acc) | _ -> mkVar hid :: acc @@ -386,10 +395,10 @@ let prove_fun_correct evd functional_induction funs_constr graphs_constr schemes (* introducing the the result of the graph and the equality hypothesis *) observe_tac "introducing" (tclMAP (fun x -> Proofview.V82.of_tactic (Simple.intro x)) [res;hres]); (* replacing [res] with its value *) - observe_tac "rewriting res value" (Proofview.V82.of_tactic (Equality.rewriteLR (EConstr.mkVar hres))); + observe_tac "rewriting res value" (Proofview.V82.of_tactic (Equality.rewriteLR (mkVar hres))); (* Conclusion *) observe_tac "exact" (fun g -> - Proofview.V82.of_tactic (exact_check (EConstr.of_constr (app_constructor g))) g) + Proofview.V82.of_tactic (exact_check (app_constructor g)) g) ] ) g @@ -401,8 +410,8 @@ let prove_fun_correct evd functional_induction funs_constr graphs_constr schemes match ctxt with | [] | [_] | [_;_] -> anomaly (Pp.str "bad context") | hres::res::decl::ctxt -> - let res = Term.it_mkLambda_or_LetIn - (Term.it_mkProd_or_LetIn concl [hres;res]) + let res = EConstr.it_mkLambda_or_LetIn + (EConstr.it_mkProd_or_LetIn concl [hres;res]) (LocalAssum (RelDecl.get_name decl, RelDecl.get_type decl) :: ctxt) in res @@ -430,7 +439,7 @@ let prove_fun_correct evd functional_induction funs_constr graphs_constr schemes List.rev (fst (List.fold_left2 (fun (bindings,avoid) decl p -> let id = Namegen.next_ident_away (Nameops.out_name (RelDecl.get_name decl)) avoid in - (EConstr.Unsafe.to_constr (nf_zeta (EConstr.of_constr p)))::bindings,id::avoid) + (nf_zeta p)::bindings,id::avoid) ([],avoid) princ_infos.predicates (lemmas))) @@ -442,7 +451,7 @@ let prove_fun_correct evd functional_induction funs_constr graphs_constr schemes observe_tac "principle" (Proofview.V82.of_tactic (assert_by (Name principle_id) princ_type - (exact_check (EConstr.of_constr f_principle)))); + (exact_check f_principle))); observe_tac "intro args_names" (tclMAP (fun id -> Proofview.V82.of_tactic (Simple.intro id)) args_names); (* observe_tac "titi" (pose_proof (Name (Id.of_string "__")) (Reductionops.nf_beta Evd.empty ((mkApp (mkVar principle_id,Array.of_list bindings))))); *) observe_tac "idtac" tclIDTAC; @@ -451,8 +460,8 @@ let prove_fun_correct evd functional_induction funs_constr graphs_constr schemes "functional_induction" ( (fun gl -> let term = mkApp (mkVar principle_id,Array.of_list bindings) in - let gl', _ty = pf_eapply (Typing.type_of ~refresh:true) gl (EConstr.of_constr term) in - Proofview.V82.of_tactic (apply (EConstr.of_constr term)) gl') + let gl', _ty = pf_eapply (Typing.type_of ~refresh:true) gl term in + Proofview.V82.of_tactic (apply term) gl') )) (fun i g -> observe_tac ("proving branche "^string_of_int i) (prove_branche i) g ) ] @@ -469,7 +478,7 @@ let generalize_dependent_of x hyp g = tclMAP (function | LocalAssum (id,t) when not (Id.equal id hyp) && - (Termops.occur_var (pf_env g) (project g) x (EConstr.of_constr t)) -> tclTHEN (Proofview.V82.of_tactic (Tactics.generalize [EConstr.mkVar id])) (thin [id]) + (Termops.occur_var (pf_env g) (project g) x (EConstr.of_constr t)) -> tclTHEN (Proofview.V82.of_tactic (Tactics.generalize [mkVar id])) (thin [id]) | _ -> tclIDTAC ) (pf_hyps g) @@ -493,44 +502,45 @@ let rec intros_with_rewrite g = and intros_with_rewrite_aux : tactic = fun g -> let eq_ind = make_eq () in - match kind_of_term (pf_concl g) with + let sigma = project g in + match EConstr.kind sigma (EConstr.of_constr (pf_concl g)) with | Prod(_,t,t') -> begin - match kind_of_term t with - | App(eq,args) when (Term.eq_constr eq eq_ind) -> - if Reductionops.is_conv (pf_env g) (project g) (EConstr.of_constr args.(1)) (EConstr.of_constr args.(2)) + match EConstr.kind sigma t with + | App(eq,args) when (EConstr.eq_constr sigma eq eq_ind) -> + if Reductionops.is_conv (pf_env g) (project g) args.(1) args.(2) then let id = pf_get_new_id (Id.of_string "y") g in tclTHENSEQ [ Proofview.V82.of_tactic (Simple.intro id); thin [id]; intros_with_rewrite ] g - else if isVar args.(1) && (Environ.evaluable_named (destVar args.(1)) (pf_env g)) + else if isVar sigma args.(1) && (Environ.evaluable_named (destVar sigma args.(1)) (pf_env g)) then tclTHENSEQ[ - Proofview.V82.of_tactic (unfold_in_concl [(Locus.AllOccurrences, Names.EvalVarRef (destVar args.(1)))]); - tclMAP (fun id -> tclTRY(Proofview.V82.of_tactic (unfold_in_hyp [(Locus.AllOccurrences, Names.EvalVarRef (destVar args.(1)))] ((destVar args.(1)),Locus.InHyp) ))) + Proofview.V82.of_tactic (unfold_in_concl [(Locus.AllOccurrences, Names.EvalVarRef (destVar sigma args.(1)))]); + tclMAP (fun id -> tclTRY(Proofview.V82.of_tactic (unfold_in_hyp [(Locus.AllOccurrences, Names.EvalVarRef (destVar sigma args.(1)))] ((destVar sigma args.(1)),Locus.InHyp) ))) (pf_ids_of_hyps g); intros_with_rewrite ] g - else if isVar args.(2) && (Environ.evaluable_named (destVar args.(2)) (pf_env g)) + else if isVar sigma args.(2) && (Environ.evaluable_named (destVar sigma args.(2)) (pf_env g)) then tclTHENSEQ[ - Proofview.V82.of_tactic (unfold_in_concl [(Locus.AllOccurrences, Names.EvalVarRef (destVar args.(2)))]); - tclMAP (fun id -> tclTRY(Proofview.V82.of_tactic (unfold_in_hyp [(Locus.AllOccurrences, Names.EvalVarRef (destVar args.(2)))] ((destVar args.(2)),Locus.InHyp) ))) + Proofview.V82.of_tactic (unfold_in_concl [(Locus.AllOccurrences, Names.EvalVarRef (destVar sigma args.(2)))]); + tclMAP (fun id -> tclTRY(Proofview.V82.of_tactic (unfold_in_hyp [(Locus.AllOccurrences, Names.EvalVarRef (destVar sigma args.(2)))] ((destVar sigma args.(2)),Locus.InHyp) ))) (pf_ids_of_hyps g); intros_with_rewrite ] g - else if isVar args.(1) + else if isVar sigma args.(1) then let id = pf_get_new_id (Id.of_string "y") g in tclTHENSEQ [ Proofview.V82.of_tactic (Simple.intro id); - generalize_dependent_of (destVar args.(1)) id; - tclTRY (Proofview.V82.of_tactic (Equality.rewriteLR (EConstr.mkVar id))); + generalize_dependent_of (destVar sigma args.(1)) id; + tclTRY (Proofview.V82.of_tactic (Equality.rewriteLR (mkVar id))); intros_with_rewrite ] g - else if isVar args.(2) + else if isVar sigma args.(2) then let id = pf_get_new_id (Id.of_string "y") g in tclTHENSEQ [ Proofview.V82.of_tactic (Simple.intro id); - generalize_dependent_of (destVar args.(2)) id; - tclTRY (Proofview.V82.of_tactic (Equality.rewriteRL (EConstr.mkVar id))); + generalize_dependent_of (destVar sigma args.(2)) id; + tclTRY (Proofview.V82.of_tactic (Equality.rewriteRL (mkVar id))); intros_with_rewrite ] g @@ -539,15 +549,15 @@ and intros_with_rewrite_aux : tactic = let id = pf_get_new_id (Id.of_string "y") g in tclTHENSEQ[ Proofview.V82.of_tactic (Simple.intro id); - tclTRY (Proofview.V82.of_tactic (Equality.rewriteLR (EConstr.mkVar id))); + tclTRY (Proofview.V82.of_tactic (Equality.rewriteLR (mkVar id))); intros_with_rewrite ] g end - | Ind _ when Term.eq_constr t (Coqlib.build_coq_False ()) -> + | Ind _ when EConstr.eq_constr sigma t (EConstr.of_constr (Coqlib.build_coq_False ())) -> Proofview.V82.of_tactic tauto g | Case(_,_,v,_) -> tclTHENSEQ[ - Proofview.V82.of_tactic (simplest_case (EConstr.of_constr v)); + Proofview.V82.of_tactic (simplest_case v); intros_with_rewrite ] g | LetIn _ -> @@ -581,10 +591,10 @@ and intros_with_rewrite_aux : tactic = let rec reflexivity_with_destruct_cases g = let destruct_case () = try - match kind_of_term (snd (destApp (pf_concl g))).(2) with + match EConstr.kind (project g) (snd (destApp (project g) (EConstr.of_constr (pf_concl g)))).(2) with | Case(_,_,v,_) -> tclTHENSEQ[ - Proofview.V82.of_tactic (simplest_case (EConstr.of_constr v)); + Proofview.V82.of_tactic (simplest_case v); Proofview.V82.of_tactic intros; observe_tac "reflexivity_with_destruct_cases" reflexivity_with_destruct_cases ] @@ -598,11 +608,11 @@ let rec reflexivity_with_destruct_cases g = match sc with None -> tclIDTAC g | Some id -> - match kind_of_term (pf_unsafe_type_of g (EConstr.mkVar id)) with - | App(eq,[|_;t1;t2|]) when Term.eq_constr eq eq_ind -> - if Equality.discriminable (pf_env g) (project g) (EConstr.of_constr t1) (EConstr.of_constr t2) + match EConstr.kind (project g) (EConstr.of_constr (pf_unsafe_type_of g (mkVar id))) with + | App(eq,[|_;t1;t2|]) when EConstr.eq_constr (project g) eq eq_ind -> + if Equality.discriminable (pf_env g) (project g) t1 t2 then Proofview.V82.of_tactic (Equality.discrHyp id) g - else if Equality.injectable (pf_env g) (project g) (EConstr.of_constr t1) (EConstr.of_constr t2) + else if Equality.injectable (pf_env g) (project g) t1 t2 then tclTHENSEQ [Proofview.V82.of_tactic (Equality.injHyp None id);thin [id];intros_with_rewrite] g else tclIDTAC g | _ -> tclIDTAC g @@ -657,7 +667,7 @@ let prove_fun_complete funcs graphs schemes lemmas_types_infos i : tactic = *) let lemmas = Array.map - (fun (_,(ctxt,concl)) -> nf_zeta (EConstr.of_constr (Termops.it_mkLambda_or_LetIn concl ctxt))) + (fun (_,(ctxt,concl)) -> nf_zeta (EConstr.it_mkLambda_or_LetIn concl ctxt)) lemmas_types_infos in (* We get the constant and the principle corresponding to this lemma *) @@ -698,7 +708,7 @@ let prove_fun_complete funcs graphs schemes lemmas_types_infos i : tactic = let rewrite_tac j ids : tactic = let graph_def = graphs.(j) in let infos = - try find_Function_infos (fst (destConst funcs.(j))) + try find_Function_infos (fst (destConst (project g) funcs.(j))) with Not_found -> error "No graph found" in if infos.is_general @@ -710,7 +720,7 @@ let prove_fun_complete funcs graphs schemes lemmas_types_infos i : tactic = in tclTHENSEQ[ tclMAP (fun id -> Proofview.V82.of_tactic (Simple.intro id)) ids; - Proofview.V82.of_tactic (Equality.rewriteLR (EConstr.of_constr (mkConst eq_lemma))); + Proofview.V82.of_tactic (Equality.rewriteLR (mkConst eq_lemma)); (* Don't forget to $\zeta$ normlize the term since the principles have been $\zeta$-normalized *) Proofview.V82.of_tactic (reduce @@ -720,11 +730,11 @@ let prove_fun_complete funcs graphs schemes lemmas_types_infos i : tactic = }) Locusops.onConcl) ; - Proofview.V82.of_tactic (generalize (List.map EConstr.mkVar ids)); + Proofview.V82.of_tactic (generalize (List.map mkVar ids)); thin ids ] else - Proofview.V82.of_tactic (unfold_in_concl [(Locus.AllOccurrences, Names.EvalConstRef (fst (destConst f)))]) + Proofview.V82.of_tactic (unfold_in_concl [(Locus.AllOccurrences, Names.EvalConstRef (fst (destConst (project g) f)))]) in (* The proof of each branche itself *) let ind_number = ref 0 in @@ -795,11 +805,10 @@ let derive_correctness make_scheme functional_induction (funs: pconstant list) ( in let type_info = (type_of_lemma_ctxt,type_of_lemma_concl) in graphs_constr.(i) <- graph; - let type_of_lemma = Term.it_mkProd_or_LetIn type_of_lemma_concl type_of_lemma_ctxt in - let _ = Typing.e_type_of (Global.env ()) evd (EConstr.of_constr type_of_lemma) in - let type_of_lemma = nf_zeta (EConstr.of_constr type_of_lemma) in - let type_of_lemma = EConstr.Unsafe.to_constr type_of_lemma in - observe (str "type_of_lemma := " ++ Printer.pr_lconstr_env (Global.env ()) !evd type_of_lemma); + let type_of_lemma = EConstr.it_mkProd_or_LetIn type_of_lemma_concl type_of_lemma_ctxt in + let _ = Typing.e_type_of (Global.env ()) evd type_of_lemma in + let type_of_lemma = nf_zeta type_of_lemma in + observe (str "type_of_lemma := " ++ Printer.pr_leconstr_env (Global.env ()) !evd type_of_lemma); type_of_lemma,type_info ) funs_constr @@ -818,7 +827,7 @@ let derive_correctness make_scheme functional_induction (funs: pconstant list) ( Array.of_list (List.map (fun entry -> - (fst (fst(Future.force entry.Entries.const_entry_body)), Option.get entry.Entries.const_entry_type ) + (EConstr.of_constr (fst (fst(Future.force entry.Entries.const_entry_body))), EConstr.of_constr (Option.get entry.Entries.const_entry_type )) ) (make_scheme evd (Array.map_to_list (fun const -> const,GType []) funs)) ) @@ -839,7 +848,7 @@ let derive_correctness make_scheme functional_induction (funs: pconstant list) ( lem_id (Decl_kinds.Global,Flags.is_universe_polymorphism (),((Decl_kinds.Proof Decl_kinds.Theorem))) !evd - typ + (EConstr.Unsafe.to_constr typ) (Lemmas.mk_hook (fun _ _ -> ())); ignore (Pfedit.by (Proofview.V82.tactic (observe_tac ("prove correctness ("^(Id.to_string f_id)^")") @@ -849,7 +858,8 @@ let derive_correctness make_scheme functional_induction (funs: pconstant list) ( (* let lem_cst = fst (destConst (Constrintern.global_reference lem_id)) in *) let _,lem_cst_constr = Evd.fresh_global (Global.env ()) !evd (Constrintern.locate_reference (Libnames.qualid_of_ident lem_id)) in - let (lem_cst,_) = destConst lem_cst_constr in + let lem_cst_constr = EConstr.of_constr lem_cst_constr in + let (lem_cst,_) = destConst !evd lem_cst_constr in update_Function {finfo with correctness_lemma = Some lem_cst}; ) @@ -863,18 +873,17 @@ let derive_correctness make_scheme functional_induction (funs: pconstant list) ( let type_info = (type_of_lemma_ctxt,type_of_lemma_concl) in graphs_constr.(i) <- graph; let type_of_lemma = - Term.it_mkProd_or_LetIn type_of_lemma_concl type_of_lemma_ctxt + EConstr.it_mkProd_or_LetIn type_of_lemma_concl type_of_lemma_ctxt in - let type_of_lemma = nf_zeta (EConstr.of_constr type_of_lemma) in - let type_of_lemma = EConstr.Unsafe.to_constr type_of_lemma in - observe (str "type_of_lemma := " ++ Printer.pr_lconstr type_of_lemma); + let type_of_lemma = nf_zeta type_of_lemma in + observe (str "type_of_lemma := " ++ Printer.pr_leconstr type_of_lemma); type_of_lemma,type_info ) funs_constr graphs_constr in - let (kn,_) as graph_ind,u = (destInd graphs_constr.(0)) in + let (kn,_) as graph_ind,u = (destInd !evd graphs_constr.(0)) in let mib,mip = Global.lookup_inductive graph_ind in let sigma, scheme = (Indrec.build_mutual_induction_scheme (Global.env ()) !evd @@ -901,7 +910,7 @@ let derive_correctness make_scheme functional_induction (funs: pconstant list) ( let lem_id = mk_complete_id f_id in Lemmas.start_proof lem_id (Decl_kinds.Global,Flags.is_universe_polymorphism (),(Decl_kinds.Proof Decl_kinds.Theorem)) sigma - (fst lemmas_types_infos.(i)) + (EConstr.Unsafe.to_constr (fst lemmas_types_infos.(i))) (Lemmas.mk_hook (fun _ _ -> ())); ignore (Pfedit.by (Proofview.V82.tactic (observe_tac ("prove completeness ("^(Id.to_string f_id)^")") @@ -910,7 +919,8 @@ let derive_correctness make_scheme functional_induction (funs: pconstant list) ( let finfo = find_Function_infos (fst f_as_constant) in let _,lem_cst_constr = Evd.fresh_global (Global.env ()) !evd (Constrintern.locate_reference (Libnames.qualid_of_ident lem_id)) in - let (lem_cst,_) = destConst lem_cst_constr in + let lem_cst_constr = EConstr.of_constr lem_cst_constr in + let (lem_cst,_) = destConst !evd lem_cst_constr in update_Function {finfo with completeness_lemma = Some lem_cst} ) funs) @@ -925,10 +935,12 @@ let derive_correctness make_scheme functional_induction (funs: pconstant list) ( if the type of hypothesis has not this form or if we cannot find the completeness lemma then we do nothing *) let revert_graph kn post_tac hid g = - let typ = pf_unsafe_type_of g (EConstr.mkVar hid) in - match kind_of_term typ with - | App(i,args) when isInd i -> - let ((kn',num) as ind'),u = destInd i in + let sigma = project g in + let typ = pf_unsafe_type_of g (mkVar hid) in + let typ = EConstr.of_constr typ in + match EConstr.kind sigma typ with + | App(i,args) when isInd sigma i -> + let ((kn',num) as ind'),u = destInd sigma i in if MutInd.equal kn kn' then (* We have generated a graph hypothesis so that we must change it if we can *) let info = @@ -945,7 +957,7 @@ let revert_graph kn post_tac hid g = let f_args,res = Array.chop (Array.length args - 1) args in tclTHENSEQ [ - Proofview.V82.of_tactic (generalize [EConstr.of_constr (applist(mkConst f_complete,(Array.to_list f_args)@[res.(0);mkVar hid]))]); + Proofview.V82.of_tactic (generalize [applist(mkConst f_complete,(Array.to_list f_args)@[res.(0);mkVar hid])]); thin [hid]; Proofview.V82.of_tactic (Simple.intro hid); post_tac hid @@ -976,20 +988,22 @@ let revert_graph kn post_tac hid g = let functional_inversion kn hid fconst f_correct : tactic = fun g -> let old_ids = List.fold_right Id.Set.add (pf_ids_of_hyps g) Id.Set.empty in - let type_of_h = pf_unsafe_type_of g (EConstr.mkVar hid) in - match kind_of_term type_of_h with - | App(eq,args) when Term.eq_constr eq (make_eq ()) -> + let sigma = project g in + let type_of_h = pf_unsafe_type_of g (mkVar hid) in + let type_of_h = EConstr.of_constr type_of_h in + match EConstr.kind sigma type_of_h with + | App(eq,args) when EConstr.eq_constr sigma eq (make_eq ()) -> let pre_tac,f_args,res = - match kind_of_term args.(1),kind_of_term args.(2) with - | App(f,f_args),_ when Term.eq_constr f fconst -> + match EConstr.kind sigma args.(1),EConstr.kind sigma args.(2) with + | App(f,f_args),_ when EConstr.eq_constr sigma f fconst -> ((fun hid -> Proofview.V82.of_tactic (intros_symmetry (Locusops.onHyp hid))),f_args,args.(2)) - |_,App(f,f_args) when Term.eq_constr f fconst -> + |_,App(f,f_args) when EConstr.eq_constr sigma f fconst -> ((fun hid -> tclIDTAC),f_args,args.(1)) | _ -> (fun hid -> tclFAIL 1 (mt ())),[||],args.(2) in tclTHENSEQ[ pre_tac hid; - Proofview.V82.of_tactic (generalize [EConstr.of_constr (applist(f_correct,(Array.to_list f_args)@[res;mkVar hid]))]); + Proofview.V82.of_tactic (generalize [applist(f_correct,(Array.to_list f_args)@[res;mkVar hid])]); thin [hid]; Proofview.V82.of_tactic (Simple.intro hid); Proofview.V82.of_tactic (Inv.inv FullInversion None (NamedHyp hid)); @@ -1028,23 +1042,25 @@ let invfun qhyp f g = Proofview.V82.of_tactic begin Tactics.try_intros_until (fun hid -> Proofview.V82.tactic begin fun g -> - let hyp_typ = pf_unsafe_type_of g (EConstr.mkVar hid) in - match kind_of_term hyp_typ with - | App(eq,args) when Term.eq_constr eq (make_eq ()) -> + let sigma = project g in + let hyp_typ = pf_unsafe_type_of g (mkVar hid) in + let hyp_typ = EConstr.of_constr hyp_typ in + match EConstr.kind sigma hyp_typ with + | App(eq,args) when EConstr.eq_constr sigma eq (make_eq ()) -> begin - let f1,_ = decompose_app args.(1) in + let f1,_ = decompose_app sigma args.(1) in try - if not (isConst f1) then failwith ""; - let finfos = find_Function_infos (fst (destConst f1)) in + if not (isConst sigma f1) then failwith ""; + let finfos = find_Function_infos (fst (destConst sigma f1)) in let f_correct = mkConst(Option.get finfos.correctness_lemma) and kn = fst finfos.graph_ind in functional_inversion kn hid f1 f_correct g with | Failure "" | Option.IsNone | Not_found -> try - let f2,_ = decompose_app args.(2) in - if not (isConst f2) then failwith ""; - let finfos = find_Function_infos (fst (destConst f2)) in + let f2,_ = decompose_app sigma args.(2) in + if not (isConst sigma f2) then failwith ""; + let finfos = find_Function_infos (fst (destConst sigma f2)) in let f_correct = mkConst(Option.get finfos.correctness_lemma) and kn = fst finfos.graph_ind in diff --git a/plugins/funind/recdef.ml b/plugins/funind/recdef.ml index 23b308efb..a80a7b5e7 100644 --- a/plugins/funind/recdef.ml +++ b/plugins/funind/recdef.ml @@ -6,7 +6,10 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +module CVars = Vars + open Term +open EConstr open Vars open Namegen open Environ @@ -42,17 +45,22 @@ open Indfun_common open Sigma.Notations open Context.Rel.Declaration +let local_assum (na, t) = + LocalAssum (na, EConstr.Unsafe.to_constr t) + +let local_def (na, b, t) = + LocalDef (na, EConstr.Unsafe.to_constr b, EConstr.Unsafe.to_constr t) (* Ugly things which should not be here *) let coq_constant m s = - Coqlib.coq_constant "RecursiveDefinition" m s + EConstr.of_constr (Coqlib.coq_constant "RecursiveDefinition" m s) let arith_Nat = ["Arith";"PeanoNat";"Nat"] let arith_Lt = ["Arith";"Lt"] let coq_init_constant s = - Coqlib.gen_constant_in_modules "RecursiveDefinition" Coqlib.init_modules s + EConstr.of_constr (Coqlib.gen_constant_in_modules "RecursiveDefinition" Coqlib.init_modules s) let find_reference sl s = let dp = Names.DirPath.make (List.rev_map Id.of_string sl) in @@ -76,8 +84,8 @@ let def_of_const t = ) |_ -> assert false -let type_of_const t = - match (kind_of_term t) with +let type_of_const sigma t = + match (EConstr.kind sigma t) with Const sp -> Typeops.type_of_constant (Global.env()) sp |_ -> assert false @@ -98,9 +106,7 @@ let nf_zeta env = let nf_betaiotazeta = (* Reductionops.local_strong Reductionops.whd_betaiotazeta *) - let clos_norm_flags flgs env sigma t = - CClosure.norm_val (CClosure.create_clos_infos flgs env) (CClosure.inject (Reductionops.nf_evar sigma t)) in - clos_norm_flags CClosure.betaiotazeta Environ.empty_env Evd.empty + Reductionops.clos_norm_flags CClosure.betaiotazeta Environ.empty_env Evd.empty @@ -116,8 +122,8 @@ let pf_get_new_ids idl g = [] let compute_renamed_type gls c = - rename_bound_vars_as_displayed (*no avoid*) [] (*no rels*) [] - (pf_unsafe_type_of gls (EConstr.of_constr c)) + EConstr.of_constr (rename_bound_vars_as_displayed (*no avoid*) [] (*no rels*) [] + (pf_unsafe_type_of gls c)) let h'_id = Id.of_string "h'" let teq_id = Id.of_string "teq" let ano_id = Id.of_string "anonymous" @@ -147,7 +153,7 @@ let coq_O = function () -> (coq_init_constant "O") let coq_S = function () -> (coq_init_constant "S") let lt_n_O = function () -> (coq_constant arith_Nat "nlt_0_r") let max_ref = function () -> (find_reference ["Recdef"] "max") -let max_constr = function () -> (constr_of_global (delayed_force max_ref)) +let max_constr = function () -> EConstr.of_constr (constr_of_global (delayed_force max_ref)) let coq_conj = function () -> find_reference Coqlib.logic_module_name "conj" let f_S t = mkApp(delayed_force coq_S, [|t|]);; @@ -166,7 +172,8 @@ let simpl_iter clause = clause (* Others ugly things ... *) -let (value_f:constr list -> global_reference -> constr) = +let (value_f:Constr.constr list -> global_reference -> Constr.constr) = + let open Term in fun al fterm -> let d0 = Loc.ghost in let rev_x_id_l = @@ -199,7 +206,7 @@ let (value_f:constr list -> global_reference -> constr) = let body = fst (understand env (Evd.from_env env) glob_body)(*FIXME*) in it_mkLambda_or_LetIn body context -let (declare_f : Id.t -> logical_kind -> constr list -> global_reference -> global_reference) = +let (declare_f : Id.t -> logical_kind -> Constr.constr list -> global_reference -> global_reference) = fun f_id kind input_type fterm_ref -> declare_fun f_id kind (value_f input_type fterm_ref);; @@ -286,7 +293,7 @@ let tclUSER tac is_mes l g = let tclUSER_if_not_mes concl_tac is_mes names_to_suppress = if is_mes - then tclCOMPLETE (fun gl -> Proofview.V82.of_tactic (Simple.apply (EConstr.of_constr (delayed_force well_founded_ltof))) gl) + then tclCOMPLETE (fun gl -> Proofview.V82.of_tactic (Simple.apply (delayed_force well_founded_ltof)) gl) else (* tclTHEN (Simple.apply (delayed_force acc_intro_generator_function) ) *) (tclUSER concl_tac is_mes names_to_suppress) @@ -301,9 +308,9 @@ let tclUSER_if_not_mes concl_tac is_mes names_to_suppress = (* [check_not_nested forbidden e] checks that [e] does not contains any variable of [forbidden] *) -let check_not_nested forbidden e = +let check_not_nested sigma forbidden e = let rec check_not_nested e = - match kind_of_term e with + match EConstr.kind sigma e with | Rel _ -> () | Var x -> if Id.List.mem x forbidden @@ -327,7 +334,7 @@ let check_not_nested forbidden e = try check_not_nested e with UserError(_,p) -> - user_err ~hdr:"_" (str "on expr : " ++ Printer.pr_lconstr e ++ str " " ++ p) + user_err ~hdr:"_" (str "on expr : " ++ Printer.pr_leconstr e ++ str " " ++ p) (* ['a info] contains the local information for traveling *) type 'a infos = @@ -374,15 +381,17 @@ type journey_info = -let rec add_vars forbidden e = - match kind_of_term e with +let add_vars sigma forbidden e = + let rec aux forbidden e = + match EConstr.kind sigma e with | Var x -> x::forbidden - | _ -> Term.fold_constr add_vars forbidden e - + | _ -> EConstr.fold sigma aux forbidden e + in + aux forbidden e let treat_case forbid_new_ids to_intros finalize_tac nb_lam e infos : tactic = fun g -> - let rev_context,b = decompose_lam_n nb_lam e in + let rev_context,b = decompose_lam_n (project g) nb_lam e in let ids = List.fold_left (fun acc (na,_) -> let pre_id = match na with @@ -402,20 +411,20 @@ let treat_case forbid_new_ids to_intros finalize_tac nb_lam e infos : tactic = Proofview.V82.of_tactic (clear to_intros); h_intros to_intros; (fun g' -> - let ty_teq = pf_unsafe_type_of g' (EConstr.mkVar heq) in + let ty_teq = pf_unsafe_type_of g' (mkVar heq) in + let ty_teq = EConstr.of_constr ty_teq in let teq_lhs,teq_rhs = - let _,args = try destApp ty_teq with DestKO -> assert false in + let _,args = try destApp (project g') ty_teq with DestKO -> assert false in args.(1),args.(2) in - let new_b' = Termops.replace_term (project g') (EConstr.of_constr teq_lhs) (EConstr.of_constr teq_rhs) (EConstr.of_constr new_b) in - let new_b' = EConstr.Unsafe.to_constr new_b' in + let new_b' = Termops.replace_term (project g') teq_lhs teq_rhs new_b in let new_infos = { infos with info = new_b'; eqs = heq::infos.eqs; forbidden_ids = if forbid_new_ids - then add_vars infos.forbidden_ids new_b' + then add_vars (project g') infos.forbidden_ids new_b' else infos.forbidden_ids } in finalize_tac new_infos g' @@ -424,8 +433,9 @@ let treat_case forbid_new_ids to_intros finalize_tac nb_lam e infos : tactic = ) ] g -let rec travel_aux jinfo continuation_tac (expr_info:constr infos) = - match kind_of_term expr_info.info with +let rec travel_aux jinfo continuation_tac (expr_info:constr infos) g = + let sigma = project g in + match EConstr.kind sigma expr_info.info with | CoFix _ | Fix _ -> error "Function cannot treat local fixpoint or cofixpoint" | Proj _ -> error "Function cannot treat projections" | LetIn(na,b,t,e) -> @@ -434,24 +444,24 @@ let rec travel_aux jinfo continuation_tac (expr_info:constr infos) = jinfo.letiN (na,b,t,e) expr_info continuation_tac in travel jinfo new_continuation_tac - {expr_info with info = b; is_final=false} + {expr_info with info = b; is_final=false} g end | Rel _ -> anomaly (Pp.str "Free var in goal conclusion !") | Prod _ -> begin try - check_not_nested (expr_info.f_id::expr_info.forbidden_ids) expr_info.info; - jinfo.otherS () expr_info continuation_tac expr_info + check_not_nested sigma (expr_info.f_id::expr_info.forbidden_ids) expr_info.info; + jinfo.otherS () expr_info continuation_tac expr_info g with e when CErrors.noncritical e -> - user_err ~hdr:"Recdef.travel" (str "the term " ++ Printer.pr_lconstr expr_info.info ++ str " can not contain a recursive call to " ++ pr_id expr_info.f_id) + user_err ~hdr:"Recdef.travel" (str "the term " ++ Printer.pr_leconstr expr_info.info ++ str " can not contain a recursive call to " ++ pr_id expr_info.f_id) end | Lambda(n,t,b) -> begin try - check_not_nested (expr_info.f_id::expr_info.forbidden_ids) expr_info.info; - jinfo.otherS () expr_info continuation_tac expr_info + check_not_nested sigma (expr_info.f_id::expr_info.forbidden_ids) expr_info.info; + jinfo.otherS () expr_info continuation_tac expr_info g with e when CErrors.noncritical e -> - user_err ~hdr:"Recdef.travel" (str "the term " ++ Printer.pr_lconstr expr_info.info ++ str " can not contain a recursive call to " ++ pr_id expr_info.f_id) + user_err ~hdr:"Recdef.travel" (str "the term " ++ Printer.pr_leconstr expr_info.info ++ str " can not contain a recursive call to " ++ pr_id expr_info.f_id) end | Case(ci,t,a,l) -> begin @@ -462,15 +472,15 @@ let rec travel_aux jinfo continuation_tac (expr_info:constr infos) = travel jinfo continuation_tac_a {expr_info with info = a; is_main_branch = false; - is_final = false} + is_final = false} g end | App _ -> - let f,args = decompose_app expr_info.info in - if Term.eq_constr f (expr_info.f_constr) - then jinfo.app_reC (f,args) expr_info continuation_tac expr_info + let f,args = decompose_app sigma expr_info.info in + if EConstr.eq_constr sigma f (expr_info.f_constr) + then jinfo.app_reC (f,args) expr_info continuation_tac expr_info g else begin - match kind_of_term f with + match EConstr.kind sigma f with | App _ -> assert false (* f is coming from a decompose_app *) | Const _ | Construct _ | Rel _ | Evar _ | Meta _ | Ind _ | Sort _ | Prod _ | Var _ -> @@ -478,15 +488,15 @@ let rec travel_aux jinfo continuation_tac (expr_info:constr infos) = let new_continuation_tac = jinfo.apP (f,args) expr_info continuation_tac in travel_args jinfo - expr_info.is_main_branch new_continuation_tac new_infos - | Case _ -> user_err ~hdr:"Recdef.travel" (str "the term " ++ Printer.pr_lconstr expr_info.info ++ str " can not contain an applied match (See Limitation in Section 2.3 of refman)") - | _ -> anomaly (Pp.str "travel_aux : unexpected "++ Printer.pr_lconstr expr_info.info) + expr_info.is_main_branch new_continuation_tac new_infos g + | Case _ -> user_err ~hdr:"Recdef.travel" (str "the term " ++ Printer.pr_leconstr expr_info.info ++ str " can not contain an applied match (See Limitation in Section 2.3 of refman)") + | _ -> anomaly (Pp.str "travel_aux : unexpected "++ Printer.pr_leconstr expr_info.info) end - | Cast(t,_,_) -> travel jinfo continuation_tac {expr_info with info=t} + | Cast(t,_,_) -> travel jinfo continuation_tac {expr_info with info=t} g | Const _ | Var _ | Meta _ | Evar _ | Sort _ | Construct _ | Ind _ -> let new_continuation_tac = jinfo.otherS () expr_info continuation_tac in - new_continuation_tac expr_info + new_continuation_tac expr_info g and travel_args jinfo is_final continuation_tac infos = let (f_args',args) = infos.info in match args with @@ -503,36 +513,37 @@ and travel_args jinfo is_final continuation_tac infos = {infos with info=arg;is_final=false} and travel jinfo continuation_tac expr_info = observe_tac - (str jinfo.message ++ Printer.pr_lconstr expr_info.info) + (str jinfo.message ++ Printer.pr_leconstr expr_info.info) (travel_aux jinfo continuation_tac expr_info) (* Termination proof *) let rec prove_lt hyple g = + let sigma = project g in begin try - let (varx,varz) = match decompose_app (pf_concl g) with - | _, x::z::_ when isVar x && isVar z -> x, z + let (varx,varz) = match decompose_app sigma (EConstr.of_constr (pf_concl g)) with + | _, x::z::_ when isVar sigma x && isVar sigma z -> x, z | _ -> assert false in let h = List.find (fun id -> - match decompose_app (pf_unsafe_type_of g (EConstr.mkVar id)) with - | _, t::_ -> Term.eq_constr t varx + match decompose_app sigma (EConstr.of_constr (pf_unsafe_type_of g (mkVar id))) with + | _, t::_ -> EConstr.eq_constr sigma t varx | _ -> false ) hyple in let y = - List.hd (List.tl (snd (decompose_app (pf_unsafe_type_of g (EConstr.mkVar h))))) in + List.hd (List.tl (snd (decompose_app sigma (EConstr.of_constr (pf_unsafe_type_of g (mkVar h)))))) in observe_tclTHENLIST (str "prove_lt1")[ - Proofview.V82.of_tactic (apply (EConstr.of_constr (mkApp(le_lt_trans (),[|varx;y;varz;mkVar h|])))); + Proofview.V82.of_tactic (apply (mkApp(le_lt_trans (),[|varx;y;varz;mkVar h|]))); observe_tac (str "prove_lt") (prove_lt hyple) ] with Not_found -> ( ( observe_tclTHENLIST (str "prove_lt2")[ - Proofview.V82.of_tactic (apply (EConstr.of_constr (delayed_force lt_S_n))); + Proofview.V82.of_tactic (apply (delayed_force lt_S_n)); (observe_tac (str "assumption: " ++ Printer.pr_goal g) (Proofview.V82.of_tactic assumption)) ]) ) @@ -550,15 +561,15 @@ let rec destruct_bounds_aux infos (bound,hyple,rechyps) lbounds g = let ids = h'::ids in let def = next_ident_away_in_goal def_id ids in observe_tclTHENLIST (str "destruct_bounds_aux1")[ - Proofview.V82.of_tactic (split (ImplicitBindings [EConstr.of_constr s_max])); + Proofview.V82.of_tactic (split (ImplicitBindings [s_max])); Proofview.V82.of_tactic (intro_then (fun id -> Proofview.V82.tactic begin observe_tac (str "destruct_bounds_aux") - (tclTHENS (Proofview.V82.of_tactic (simplest_case (EConstr.mkVar id))) + (tclTHENS (Proofview.V82.of_tactic (simplest_case (mkVar id))) [ observe_tclTHENLIST (str "")[Proofview.V82.of_tactic (intro_using h_id); - Proofview.V82.of_tactic (simplest_elim(EConstr.of_constr (mkApp(delayed_force lt_n_O,[|s_max|])))); + Proofview.V82.of_tactic (simplest_elim(mkApp(delayed_force lt_n_O,[|s_max|]))); Proofview.V82.of_tactic default_full_auto]; observe_tclTHENLIST (str "destruct_bounds_aux2")[ observe_tac (str "clearing k ") (Proofview.V82.of_tactic (clear [id])); @@ -589,7 +600,7 @@ let rec destruct_bounds_aux infos (bound,hyple,rechyps) lbounds g = ] g | (_,v_bound)::l -> observe_tclTHENLIST (str "destruct_bounds_aux3")[ - Proofview.V82.of_tactic (simplest_elim (EConstr.mkVar v_bound)); + Proofview.V82.of_tactic (simplest_elim (mkVar v_bound)); Proofview.V82.of_tactic (clear [v_bound]); tclDO 2 (Proofview.V82.of_tactic intro); onNthHypId 1 @@ -598,7 +609,7 @@ let rec destruct_bounds_aux infos (bound,hyple,rechyps) lbounds g = (fun p -> observe_tclTHENLIST (str "destruct_bounds_aux4")[ Proofview.V82.of_tactic (simplest_elim - (EConstr.of_constr (mkApp(delayed_force max_constr, [| bound; mkVar p|])))); + (mkApp(delayed_force max_constr, [| bound; mkVar p|]))); tclDO 3 (Proofview.V82.of_tactic intro); onNLastHypsId 3 (fun lids -> match lids with @@ -623,7 +634,7 @@ let terminate_app f_and_args expr_info continuation_tac infos = observe_tclTHENLIST (str "terminate_app1")[ continuation_tac infos; observe_tac (str "first split") - (Proofview.V82.of_tactic (split (ImplicitBindings [EConstr.of_constr infos.info]))); + (Proofview.V82.of_tactic (split (ImplicitBindings [infos.info]))); observe_tac (str "destruct_bounds (1)") (destruct_bounds infos) ] else continuation_tac infos @@ -634,17 +645,18 @@ let terminate_others _ expr_info continuation_tac infos = observe_tclTHENLIST (str "terminate_others")[ continuation_tac infos; observe_tac (str "first split") - (Proofview.V82.of_tactic (split (ImplicitBindings [EConstr.of_constr infos.info]))); + (Proofview.V82.of_tactic (split (ImplicitBindings [infos.info]))); observe_tac (str "destruct_bounds") (destruct_bounds infos) ] else continuation_tac infos -let terminate_letin (na,b,t,e) expr_info continuation_tac info = +let terminate_letin (na,b,t,e) expr_info continuation_tac info g = + let sigma = project g in let new_e = subst1 info.info e in let new_forbidden = let forbid = try - check_not_nested (expr_info.f_id::expr_info.forbidden_ids) b; + check_not_nested sigma (expr_info.f_id::expr_info.forbidden_ids) b; true with e when CErrors.noncritical e -> false in @@ -655,7 +667,7 @@ let terminate_letin (na,b,t,e) expr_info continuation_tac info = | Name id -> id::info.forbidden_ids else info.forbidden_ids in - continuation_tac {info with info = new_e; forbidden_ids = new_forbidden} + continuation_tac {info with info = new_e; forbidden_ids = new_forbidden} g let pf_type c tac gl = let evars, ty = Typing.type_of (pf_env gl) (project gl) c in @@ -682,30 +694,31 @@ let mkDestructEq : (fun decl -> let open Context.Named.Declaration in let id = get_id decl in - if Id.List.mem id not_on_hyp || not (Termops.occur_term (project g) (EConstr.of_constr expr) (EConstr.of_constr (get_type decl))) + if Id.List.mem id not_on_hyp || not (Termops.occur_term (project g) expr (EConstr.of_constr (get_type decl))) then None else Some id) hyps in let to_revert_constr = List.rev_map mkVar to_revert in - let type_of_expr = pf_unsafe_type_of g (EConstr.of_constr expr) in + let type_of_expr = pf_unsafe_type_of g expr in + let type_of_expr = EConstr.of_constr type_of_expr in let new_hyps = mkApp(Lazy.force refl_equal, [|type_of_expr; expr|]):: to_revert_constr in - let new_hyps = List.map EConstr.of_constr new_hyps in pf_typel new_hyps (fun _ -> observe_tclTHENLIST (str "mkDestructEq") [Proofview.V82.of_tactic (generalize new_hyps); (fun g2 -> let changefun patvars = { run = fun sigma -> - let redfun = pattern_occs [Locus.AllOccurrencesBut [1], EConstr.of_constr expr] in + let redfun = pattern_occs [Locus.AllOccurrencesBut [1], expr] in let Sigma (c, sigma, p) = redfun.Reductionops.e_redfun (pf_env g2) sigma (EConstr.of_constr (pf_concl g2)) in Sigma (c, sigma, p) } in Proofview.V82.of_tactic (change_in_concl None changefun) g2); - Proofview.V82.of_tactic (simplest_case (EConstr.of_constr expr))]), to_revert + Proofview.V82.of_tactic (simplest_case expr)]), to_revert let terminate_case next_step (ci,a,t,l) expr_info continuation_tac infos g = + let sigma = project g in let f_is_present = try - check_not_nested (expr_info.f_id::expr_info.forbidden_ids) a; + check_not_nested sigma (expr_info.f_id::expr_info.forbidden_ids) a; false with e when CErrors.noncritical e -> true @@ -719,7 +732,7 @@ let terminate_case next_step (ci,a,t,l) expr_info continuation_tac infos g = let destruct_tac,rev_to_thin_intro = mkDestructEq [expr_info.rec_arg_id] a' g in let to_thin_intro = List.rev rev_to_thin_intro in - observe_tac (str "treating cases (" ++ int (Array.length l) ++ str")" ++ spc () ++ Printer.pr_lconstr a') + observe_tac (str "treating cases (" ++ int (Array.length l) ++ str")" ++ spc () ++ Printer.pr_leconstr a') (try (tclTHENS destruct_tac @@ -728,16 +741,17 @@ let terminate_case next_step (ci,a,t,l) expr_info continuation_tac infos g = with | UserError(Some "Refiner.thensn_tac3",_) | UserError(Some "Refiner.tclFAIL_s",_) -> - (observe_tac (str "is computable " ++ Printer.pr_lconstr new_info.info) (next_step continuation_tac {new_info with info = nf_betaiotazeta new_info.info} ) + (observe_tac (str "is computable " ++ Printer.pr_leconstr new_info.info) (next_step continuation_tac {new_info with info = nf_betaiotazeta new_info.info} ) )) g -let terminate_app_rec (f,args) expr_info continuation_tac _ = - List.iter (check_not_nested (expr_info.f_id::expr_info.forbidden_ids)) +let terminate_app_rec (f,args) expr_info continuation_tac _ g = + let sigma = project g in + List.iter (check_not_nested sigma (expr_info.f_id::expr_info.forbidden_ids)) args; begin try - let v = List.assoc_f (List.equal Constr.equal) args expr_info.args_assoc in + let v = List.assoc_f (List.equal (EConstr.eq_constr sigma)) args expr_info.args_assoc in let new_infos = {expr_info with info = v} in observe_tclTHENLIST (str "terminate_app_rec")[ continuation_tac new_infos; @@ -745,16 +759,16 @@ let terminate_app_rec (f,args) expr_info continuation_tac _ = then observe_tclTHENLIST (str "terminate_app_rec1")[ observe_tac (str "first split") - (Proofview.V82.of_tactic (split (ImplicitBindings [EConstr.of_constr new_infos.info]))); + (Proofview.V82.of_tactic (split (ImplicitBindings [new_infos.info]))); observe_tac (str "destruct_bounds (3)") (destruct_bounds new_infos) ] else tclIDTAC - ] + ] g with Not_found -> observe_tac (str "terminate_app_rec not found") (tclTHENS - (Proofview.V82.of_tactic (simplest_elim (EConstr.of_constr (mkApp(mkVar expr_info.ih,Array.of_list args))))) + (Proofview.V82.of_tactic (simplest_elim (mkApp(mkVar expr_info.ih,Array.of_list args)))) [ observe_tclTHENLIST (str "terminate_app_rec2")[ Proofview.V82.of_tactic (intro_using rec_res_id); @@ -775,7 +789,7 @@ let terminate_app_rec (f,args) expr_info continuation_tac _ = then observe_tclTHENLIST (str "terminate_app_rec4")[ observe_tac (str "first split") - (Proofview.V82.of_tactic (split (ImplicitBindings [EConstr.of_constr new_infos.info]))); + (Proofview.V82.of_tactic (split (ImplicitBindings [new_infos.info]))); observe_tac (str "destruct_bounds (2)") (destruct_bounds new_infos) ] @@ -788,7 +802,7 @@ let terminate_app_rec (f,args) expr_info continuation_tac _ = ]; observe_tac (str "proving decreasing") ( tclTHENS (* proof of args < formal args *) - (Proofview.V82.of_tactic (apply (EConstr.of_constr (Lazy.force expr_info.acc_inv)))) + (Proofview.V82.of_tactic (apply (Lazy.force expr_info.acc_inv))) [ observe_tac (str "assumption") (Proofview.V82.of_tactic assumption); observe_tclTHENLIST (str "terminate_app_rec5") @@ -808,7 +822,7 @@ let terminate_app_rec (f,args) expr_info continuation_tac _ = ); ] ]) - ]) + ]) g end let terminate_info = @@ -830,26 +844,28 @@ let equation_case next_step (ci,a,t,l) expr_info continuation_tac infos = observe_tac (str "equation case") (terminate_case next_step (ci,a,t,l) expr_info continuation_tac infos) let rec prove_le g = + let sigma = project g in let x,z = - let _,args = decompose_app (pf_concl g) in + let _,args = decompose_app sigma (EConstr.of_constr (pf_concl g)) in (List.hd args,List.hd (List.tl args)) in tclFIRST[ Proofview.V82.of_tactic assumption; - Proofview.V82.of_tactic (apply (EConstr.of_constr (delayed_force le_n))); + Proofview.V82.of_tactic (apply (delayed_force le_n)); begin try let matching_fun = pf_is_matching g - (Pattern.PApp(Pattern.PRef (reference_of_constr (le ())),[|Pattern.PVar (destVar x);Pattern.PMeta None|])) in + (Pattern.PApp(Pattern.PRef (reference_of_constr (EConstr.Unsafe.to_constr (le ()))),[|Pattern.PVar (destVar sigma x);Pattern.PMeta None|])) in let (h,t) = List.find (fun (_,t) -> matching_fun (EConstr.of_constr t)) (pf_hyps_types g) in + let t = EConstr.of_constr t in let y = - let _,args = decompose_app t in + let _,args = decompose_app sigma t in List.hd (List.tl args) in observe_tclTHENLIST (str "prove_le")[ - Proofview.V82.of_tactic (apply(EConstr.of_constr (mkApp(le_trans (),[|x;y;z;mkVar h|])))); + Proofview.V82.of_tactic (apply(mkApp(le_trans (),[|x;y;z;mkVar h|]))); observe_tac (str "prove_le (rec)") (prove_le) ] with Not_found -> tclFAIL 0 (mt()) @@ -863,23 +879,24 @@ let rec make_rewrite_list expr_info max = function observe_tac (str "make_rewrite_list") (tclTHENS (observe_tac (str "rewrite heq on " ++ pr_id p ) ( (fun g -> + let sigma = project g in let t_eq = compute_renamed_type g (mkVar hp) in let k,def = - let k_na,_,t = destProd t_eq in - let _,_,t = destProd t in - let def_na,_,_ = destProd t in + let k_na,_,t = destProd sigma t_eq in + let _,_,t = destProd sigma t in + let def_na,_,_ = destProd sigma t in Nameops.out_name k_na,Nameops.out_name def_na in Proofview.V82.of_tactic (general_rewrite_bindings false Locus.AllOccurrences true (* dep proofs also: *) true - (EConstr.mkVar hp, + (mkVar hp, ExplicitBindings[Loc.ghost,NamedHyp def, - EConstr.of_constr expr_info.f_constr;Loc.ghost,NamedHyp k, - EConstr.of_constr (f_S max)]) false) g) ) + expr_info.f_constr;Loc.ghost,NamedHyp k, + f_S max]) false) g) ) ) [make_rewrite_list expr_info max l; observe_tclTHENLIST (str "make_rewrite_list")[ (* x < S max proof *) - Proofview.V82.of_tactic (apply (EConstr.of_constr (delayed_force le_lt_n_Sm))); + Proofview.V82.of_tactic (apply (delayed_force le_lt_n_Sm)); observe_tac (str "prove_le(2)") prove_le ] ] ) @@ -889,20 +906,21 @@ let make_rewrite expr_info l hp max = (observe_tac (str "make_rewrite") (make_rewrite_list expr_info max l)) (observe_tac (str "make_rewrite") (tclTHENS (fun g -> + let sigma = project g in let t_eq = compute_renamed_type g (mkVar hp) in let k,def = - let k_na,_,t = destProd t_eq in - let _,_,t = destProd t in - let def_na,_,_ = destProd t in + let k_na,_,t = destProd sigma t_eq in + let _,_,t = destProd sigma t in + let def_na,_,_ = destProd sigma t in Nameops.out_name k_na,Nameops.out_name def_na in observe_tac (str "general_rewrite_bindings") (Proofview.V82.of_tactic (general_rewrite_bindings false Locus.AllOccurrences true (* dep proofs also: *) true - (EConstr.mkVar hp, + (mkVar hp, ExplicitBindings[Loc.ghost,NamedHyp def, - EConstr.of_constr expr_info.f_constr;Loc.ghost,NamedHyp k, - EConstr.of_constr (f_S (f_S max))]) false)) g) + expr_info.f_constr;Loc.ghost,NamedHyp k, + f_S (f_S max)]) false)) g) [observe_tac(str "make_rewrite finalize") ( (* tclORELSE( h_reflexivity) *) (observe_tclTHENLIST (str "make_rewrite")[ @@ -931,7 +949,7 @@ let rec compute_max rew_tac max l = | (_,p,_)::l -> observe_tclTHENLIST (str "compute_max")[ Proofview.V82.of_tactic (simplest_elim - (EConstr.of_constr (mkApp(delayed_force max_constr, [| max; mkVar p|])))); + (mkApp(delayed_force max_constr, [| max; mkVar p|]))); tclDO 3 (Proofview.V82.of_tactic intro); onNLastHypsId 3 (fun lids -> match lids with @@ -950,7 +968,7 @@ let rec destruct_hex expr_info acc l = end | (v,hex)::l -> observe_tclTHENLIST (str "destruct_hex")[ - Proofview.V82.of_tactic (simplest_case (EConstr.mkVar hex)); + Proofview.V82.of_tactic (simplest_case (mkVar hex)); Proofview.V82.of_tactic (clear [hex]); tclDO 2 (Proofview.V82.of_tactic intro); onNthHypId 1 (fun hp -> @@ -977,36 +995,37 @@ let rec intros_values_eq expr_info acc = let equation_others _ expr_info continuation_tac infos = if expr_info.is_final && expr_info.is_main_branch then - observe_tac (str "equation_others (cont_tac +intros) " ++ Printer.pr_lconstr expr_info.info) + observe_tac (str "equation_others (cont_tac +intros) " ++ Printer.pr_leconstr expr_info.info) (tclTHEN (continuation_tac infos) - (observe_tac (str "intros_values_eq equation_others " ++ Printer.pr_lconstr expr_info.info) (intros_values_eq expr_info []))) - else observe_tac (str "equation_others (cont_tac) " ++ Printer.pr_lconstr expr_info.info) (continuation_tac infos) + (observe_tac (str "intros_values_eq equation_others " ++ Printer.pr_leconstr expr_info.info) (intros_values_eq expr_info []))) + else observe_tac (str "equation_others (cont_tac) " ++ Printer.pr_leconstr expr_info.info) (continuation_tac infos) let equation_app f_and_args expr_info continuation_tac infos = if expr_info.is_final && expr_info.is_main_branch then ((observe_tac (str "intros_values_eq equation_app") (intros_values_eq expr_info []))) else continuation_tac infos -let equation_app_rec (f,args) expr_info continuation_tac info = +let equation_app_rec (f,args) expr_info continuation_tac info g = + let sigma = project g in begin try - let v = List.assoc_f (List.equal Constr.equal) args expr_info.args_assoc in + let v = List.assoc_f (List.equal (EConstr.eq_constr sigma)) args expr_info.args_assoc in let new_infos = {expr_info with info = v} in - observe_tac (str "app_rec found") (continuation_tac new_infos) + observe_tac (str "app_rec found") (continuation_tac new_infos) g with Not_found -> if expr_info.is_final && expr_info.is_main_branch then observe_tclTHENLIST (str "equation_app_rec") - [ Proofview.V82.of_tactic (simplest_case (EConstr.of_constr (mkApp (expr_info.f_terminate,Array.of_list args)))); + [ Proofview.V82.of_tactic (simplest_case (mkApp (expr_info.f_terminate,Array.of_list args))); continuation_tac {expr_info with args_assoc = (args,delayed_force coq_O)::expr_info.args_assoc}; observe_tac (str "app_rec intros_values_eq") (intros_values_eq expr_info []) - ] + ] g else observe_tclTHENLIST (str "equation_app_rec1")[ - Proofview.V82.of_tactic (simplest_case (EConstr.of_constr (mkApp (expr_info.f_terminate,Array.of_list args)))); + Proofview.V82.of_tactic (simplest_case (mkApp (expr_info.f_terminate,Array.of_list args))); observe_tac (str "app_rec not_found") (continuation_tac {expr_info with args_assoc = (args,delayed_force coq_O)::expr_info.args_assoc}) - ] + ] g end let equation_info = @@ -1025,6 +1044,8 @@ let prove_eq = travel equation_info (* [compute_terminate_type] computes the type of the Definition f_terminate from the type of f_F *) let compute_terminate_type nb_args func = + let open Term in + let open CVars in let _,a_arrow_b,_ = destLambda(def_of_const (constr_of_global func)) in let rev_args,b = decompose_prod_n nb_args a_arrow_b in let left = @@ -1037,6 +1058,7 @@ let compute_terminate_type nb_args func = ) in let right = mkRel 5 in + let delayed_force c = EConstr.Unsafe.to_constr (delayed_force c) in let equality = mkApp(delayed_force eq, [|lift 5 b; left; right|]) in let result = (mkProd ((Name def_id) , lift 4 a_arrow_b, equality)) in let cond = mkApp(delayed_force lt, [|(mkRel 2); (mkRel 1)|]) in @@ -1049,7 +1071,7 @@ let compute_terminate_type nb_args func = delayed_force nat, (mkProd (Name k_id, delayed_force nat, mkArrow cond result))))|])in - let value = mkApp(constr_of_global (delayed_force coq_sig_ref), + let value = mkApp(constr_of_global (Util.delayed_force coq_sig_ref), [|b; (mkLambda (Name v_id, b, nb_iter))|]) in compose_prod rev_args value @@ -1089,9 +1111,9 @@ let termination_proof_header is_mes input_type ids args_id relation (str "first assert") (Proofview.V82.of_tactic (assert_before (Name wf_rec_arg) - (EConstr.of_constr (mkApp (delayed_force acc_rel, + (mkApp (delayed_force acc_rel, [|input_type;relation;mkVar rec_arg_id|]) - )) + ) )) ) [ @@ -1101,7 +1123,7 @@ let termination_proof_header is_mes input_type ids args_id relation (str "second assert") (Proofview.V82.of_tactic (assert_before (Name wf_thm) - (EConstr.of_constr (mkApp (delayed_force well_founded,[|input_type;relation|]))) + (mkApp (delayed_force well_founded,[|input_type;relation|])) )) ) [ @@ -1110,7 +1132,7 @@ let termination_proof_header is_mes input_type ids args_id relation (* this gives the accessibility argument *) observe_tac (str "apply wf_thm") - (Proofview.V82.of_tactic (Simple.apply (EConstr.of_constr (mkApp(mkVar wf_thm,[|mkVar rec_arg_id|])))) + (Proofview.V82.of_tactic (Simple.apply (mkApp(mkVar wf_thm,[|mkVar rec_arg_id|]))) ) ] ; @@ -1119,7 +1141,7 @@ let termination_proof_header is_mes input_type ids args_id relation [observe_tac (str "generalize") (onNLastHypsId (nargs+1) (tclMAP (fun id -> - tclTHEN (Proofview.V82.of_tactic (Tactics.generalize [EConstr.mkVar id])) (Proofview.V82.of_tactic (clear [id]))) + tclTHEN (Proofview.V82.of_tactic (Tactics.generalize [mkVar id])) (Proofview.V82.of_tactic (clear [id]))) )) ; observe_tac (str "fix") (Proofview.V82.of_tactic (fix (Some hrec) (nargs+1))); @@ -1133,25 +1155,27 @@ let termination_proof_header is_mes input_type ids args_id relation -let rec instantiate_lambda t l = +let rec instantiate_lambda sigma t l = match l with | [] -> t | a::l -> - let (_, _, body) = destLambda t in - instantiate_lambda (subst1 a body) l + let (_, _, body) = destLambda sigma t in + instantiate_lambda sigma (subst1 a body) l let whole_start (concl_tac:tactic) nb_args is_mes func input_type relation rec_arg_num : tactic = begin fun g -> + let sigma = project g in let ids = Termops.ids_of_named_context (pf_hyps g) in let func_body = (def_of_const (constr_of_global func)) in - let (f_name, _, body1) = destLambda func_body in + let func_body = EConstr.of_constr func_body in + let (f_name, _, body1) = destLambda sigma func_body in let f_id = match f_name with | Name f_id -> next_ident_away_in_goal f_id ids | Anonymous -> anomaly (Pp.str "Anonymous function") in - let n_names_types,_ = decompose_lam_n nb_args body1 in + let n_names_types,_ = decompose_lam_n sigma nb_args body1 in let n_ids,ids = List.fold_left (fun (n_ids,ids) (n_name,_) -> @@ -1165,7 +1189,7 @@ let whole_start (concl_tac:tactic) nb_args is_mes func input_type relation rec_a n_names_types in let rec_arg_id = List.nth n_ids (rec_arg_num - 1) in - let expr = instantiate_lambda func_body (mkVar f_id::(List.map mkVar n_ids)) in + let expr = instantiate_lambda sigma func_body (mkVar f_id::(List.map mkVar n_ids)) in termination_proof_header is_mes input_type @@ -1207,17 +1231,17 @@ let get_current_subgoals_types () = let { Evd.it=sgs ; sigma=sigma } = Proof.V82.subgoals p in sigma, List.map (Goal.V82.abstract_type sigma) sgs -let build_and_l l = +let build_and_l sigma l = let and_constr = Coqlib.build_coq_and () in let conj_constr = coq_conj () in let mk_and p1 p2 = - Term.mkApp(and_constr,[|p1;p2|]) in + mkApp(EConstr.of_constr and_constr,[|p1;p2|]) in let rec is_well_founded t = - match kind_of_term t with + match EConstr.kind sigma t with | Prod(_,_,t') -> is_well_founded t' | App(_,_) -> - let (f,_) = decompose_app t in - Term.eq_constr f (well_founded ()) + let (f,_) = decompose_app sigma t in + EConstr.eq_constr sigma f (well_founded ()) | _ -> false in @@ -1248,16 +1272,16 @@ let is_rec_res id = String.equal (String.sub id_name 0 (String.length rec_res_name)) rec_res_name with Invalid_argument _ -> false -let clear_goals = +let clear_goals sigma = let rec clear_goal t = - match kind_of_term t with + match EConstr.kind sigma t with | Prod(Name id as na,t',b) -> let b' = clear_goal b in - if noccurn 1 b' && (is_rec_res id) + if noccurn sigma 1 b' && (is_rec_res id) then Vars.lift (-1) b' else if b' == b then t else mkProd(na,t',b') - | _ -> Term.map_constr clear_goal t + | _ -> EConstr.map sigma clear_goal t in List.map clear_goal @@ -1265,9 +1289,9 @@ let clear_goals = let build_new_goal_type () = let sigma, sub_gls_types = get_current_subgoals_types () in (* Pp.msgnl (str "sub_gls_types1 := " ++ Util.prlist_with_sep (fun () -> Pp.fnl () ++ Pp.fnl ()) Printer.pr_lconstr sub_gls_types); *) - let sub_gls_types = clear_goals sub_gls_types in + let sub_gls_types = clear_goals sigma sub_gls_types in (* Pp.msgnl (str "sub_gls_types2 := " ++ Pp.prlist_with_sep (fun () -> Pp.fnl () ++ Pp.fnl ()) Printer.pr_lconstr sub_gls_types); *) - let res = build_and_l sub_gls_types in + let res = build_and_l sigma sub_gls_types in sigma, res let is_opaque_constant c = @@ -1288,7 +1312,7 @@ let open_new_goal build_proof sigma using_lemmas ref_ goal_name (gls_type,decomp anomaly (Pp.str "open_new_goal with an unamed theorem") in let na = next_global_ident_away name [] in - if Termops.occur_existential sigma (EConstr.of_constr gls_type) then + if Termops.occur_existential sigma gls_type then CErrors.error "\"abstract\" cannot handle existentials"; let hook _ _ = let opacity = @@ -1299,8 +1323,7 @@ let open_new_goal build_proof sigma using_lemmas ref_ goal_name (gls_type,decomp | _ -> anomaly ~label:"equation_lemma" (Pp.str "not a constant") in let lemma = mkConst (Names.Constant.make1 (Lib.make_kn na)) in - ref_ := Some lemma ; - let lemma = EConstr.of_constr lemma in + ref_ := Some (EConstr.Unsafe.to_constr lemma); let lid = ref [] in let h_num = ref (-1) in let env = Global.env () in @@ -1315,7 +1338,7 @@ let open_new_goal build_proof sigma using_lemmas ref_ goal_name (gls_type,decomp (fun g -> let ids = pf_ids_of_hyps g in tclTHEN - (Proofview.V82.of_tactic (Elim.h_decompose_and (EConstr.mkVar hid))) + (Proofview.V82.of_tactic (Elim.h_decompose_and (mkVar hid))) (fun g -> let ids' = pf_ids_of_hyps g in lid := List.rev (List.subtract Id.equal ids' ids); @@ -1326,8 +1349,9 @@ let open_new_goal build_proof sigma using_lemmas ref_ goal_name (gls_type,decomp ); ] gls) (fun g -> - match kind_of_term (pf_concl g) with - | App(f,_) when Term.eq_constr f (well_founded ()) -> + let sigma = project g in + match EConstr.kind sigma (EConstr.of_constr (pf_concl g)) with + | App(f,_) when EConstr.eq_constr sigma f (well_founded ()) -> Proofview.V82.of_tactic (Auto.h_auto None [] (Some [])) g | _ -> incr h_num; @@ -1336,11 +1360,11 @@ let open_new_goal build_proof sigma using_lemmas ref_ goal_name (gls_type,decomp tclCOMPLETE( tclFIRST[ tclTHEN - (Proofview.V82.of_tactic (eapply_with_bindings (EConstr.mkVar (List.nth !lid !h_num), NoBindings))) + (Proofview.V82.of_tactic (eapply_with_bindings (mkVar (List.nth !lid !h_num), NoBindings))) (Proofview.V82.of_tactic e_assumption); Eauto.eauto_with_bases (true,5) - [{ Tacexpr.delayed = fun _ sigma -> Sigma.here (EConstr.of_constr (Lazy.force refl_equal)) sigma}] + [{ Tacexpr.delayed = fun _ sigma -> Sigma.here (Lazy.force refl_equal) sigma}] [Hints.Hint_db.empty empty_transparent_state false] ] ) @@ -1353,7 +1377,7 @@ let open_new_goal build_proof sigma using_lemmas ref_ goal_name (gls_type,decomp Lemmas.start_proof na (Decl_kinds.Global, false (* FIXME *), Decl_kinds.Proof Decl_kinds.Lemma) - sigma gls_type + sigma (EConstr.Unsafe.to_constr gls_type) (Lemmas.mk_hook hook); if Indfun_common.is_strict_tcc () then @@ -1387,7 +1411,7 @@ let open_new_goal build_proof sigma using_lemmas ref_ goal_name (gls_type,decomp let com_terminate tcc_lemma_name - tcc_lemma_ref + (tcc_lemma_ref : Constr.t option ref) is_mes fonctional_ref input_type @@ -1424,22 +1448,25 @@ let com_terminate let start_equation (f:global_reference) (term_f:global_reference) (cont_tactic:Id.t list -> tactic) g = + let sigma = project g in let ids = pf_ids_of_hyps g in let terminate_constr = constr_of_global term_f in - let nargs = nb_prod (project g) (EConstr.of_constr (fst (type_of_const terminate_constr))) (*FIXME*) in + let terminate_constr = EConstr.of_constr terminate_constr in + let nargs = nb_prod (project g) (EConstr.of_constr (fst (type_of_const sigma terminate_constr))) (*FIXME*) in let x = n_x_id ids nargs in observe_tac (str "start_equation") (observe_tclTHENLIST (str "start_equation") [ h_intros x; Proofview.V82.of_tactic (unfold_in_concl [(Locus.AllOccurrences, evaluable_of_global_reference f)]); observe_tac (str "simplest_case") - (Proofview.V82.of_tactic (simplest_case (EConstr.of_constr (mkApp (terminate_constr, - Array.of_list (List.map mkVar x)))))); + (Proofview.V82.of_tactic (simplest_case (mkApp (terminate_constr, + Array.of_list (List.map mkVar x))))); observe_tac (str "prove_eq") (cont_tactic x)]) g;; let (com_eqn : int -> Id.t -> global_reference -> global_reference -> global_reference - -> constr -> unit) = + -> Constr.constr -> unit) = fun nb_arg eq_name functional_ref f_ref terminate_ref equation_lemma_type -> + let open CVars in let opacity = match terminate_ref with | ConstRef c -> is_opaque_constant c @@ -1459,13 +1486,13 @@ let (com_eqn : int -> Id.t -> (fun x -> prove_eq (fun _ -> tclIDTAC) {nb_arg=nb_arg; - f_terminate = constr_of_global terminate_ref; - f_constr = f_constr; + f_terminate = EConstr.of_constr (constr_of_global terminate_ref); + f_constr = EConstr.of_constr f_constr; concl_tac = tclIDTAC; func=functional_ref; - info=(instantiate_lambda - (def_of_const (constr_of_global functional_ref)) - (f_constr::List.map mkVar x) + info=(instantiate_lambda Evd.empty + (EConstr.of_constr (def_of_const (constr_of_global functional_ref))) + (EConstr.of_constr f_constr::List.map mkVar x) ); is_main_branch = true; is_final = true; @@ -1491,6 +1518,8 @@ let (com_eqn : int -> Id.t -> let recursive_definition is_mes function_name rec_impls type_of_f r rec_arg_num eq generate_induction_principle using_lemmas : unit = + let open Term in + let open CVars in let env = Global.env() in let evd = ref (Evd.from_env env) in let function_type = interp_type_evars env evd type_of_f in @@ -1498,8 +1527,9 @@ let recursive_definition is_mes function_name rec_impls type_of_f r rec_arg_num (* Pp.msgnl (str "function type := " ++ Printer.pr_lconstr function_type); *) let ty = interp_type_evars env evd ~impls:rec_impls eq in let evm, nf = Evarutil.nf_evars_and_universes !evd in - let equation_lemma_type = nf_betaiotazeta (nf ty) in + let equation_lemma_type = nf_betaiotazeta (EConstr.of_constr (nf ty)) in let function_type = nf function_type in + let equation_lemma_type = EConstr.Unsafe.to_constr equation_lemma_type in (* Pp.msgnl (str "lemma type := " ++ Printer.pr_lconstr equation_lemma_type ++ fnl ()); *) let res_vars,eq' = decompose_prod equation_lemma_type in let env_eq' = Environ.push_rel_context (List.map (fun (x,y) -> LocalAssum (x,y)) res_vars) env in @@ -1557,7 +1587,7 @@ let recursive_definition is_mes function_name rec_impls type_of_f r rec_arg_num and functional_ref = destConst (constr_of_global functional_ref) and eq_ref = destConst (constr_of_global eq_ref) in generate_induction_principle f_ref tcc_lemma_constr - functional_ref eq_ref rec_arg_num rec_arg_type (nb_prod evm (EConstr.of_constr res)) relation; + functional_ref eq_ref rec_arg_num (EConstr.of_constr rec_arg_type) (nb_prod evm (EConstr.of_constr res)) (EConstr.of_constr relation); if Flags.is_verbose () then msgnl (h 1 (Ppconstr.pr_id function_name ++ spc () ++ str"is defined" )++ fnl () ++ @@ -1570,8 +1600,8 @@ let recursive_definition is_mes function_name rec_impls type_of_f r rec_arg_num tcc_lemma_name tcc_lemma_constr is_mes functional_ref - rec_arg_type - relation rec_arg_num + (EConstr.of_constr rec_arg_type) + (EConstr.of_constr relation) rec_arg_num term_id using_lemmas (List.length res_vars) diff --git a/plugins/funind/recdef.mli b/plugins/funind/recdef.mli index f60eedbe6..9c1081b9d 100644 --- a/plugins/funind/recdef.mli +++ b/plugins/funind/recdef.mli @@ -15,6 +15,6 @@ bool -> int -> Constrexpr.constr_expr -> (Term.pconstant -> Term.constr option ref -> Term.pconstant -> - Term.pconstant -> int -> Term.types -> int -> Term.constr -> 'a) -> Constrexpr.constr_expr list -> unit + Term.pconstant -> int -> EConstr.types -> int -> EConstr.constr -> 'a) -> Constrexpr.constr_expr list -> unit diff --git a/proofs/goal.ml b/proofs/goal.ml index 7499bc251..410e738de 100644 --- a/proofs/goal.ml +++ b/proofs/goal.ml @@ -49,7 +49,7 @@ module V82 = struct (* Access to ".evar_concl" *) let concl evars gl = let evi = Evd.find evars gl in - evi.Evd.evar_concl + EConstr.of_constr evi.Evd.evar_concl (* Access to ".evar_extra" *) let extra evars gl = @@ -146,6 +146,7 @@ module V82 = struct (* Goal represented as a type, doesn't take into account section variables *) let abstract_type sigma gl = + let open EConstr in let (gl,sigma) = nf_evar sigma gl in let env = env sigma gl in let genv = Global.env () in @@ -159,6 +160,4 @@ module V82 = struct t ) ~init:(concl sigma gl) env - let concl sigma gl = EConstr.of_constr (concl sigma gl) - end diff --git a/proofs/goal.mli b/proofs/goal.mli index ca6584e76..a2fa34c05 100644 --- a/proofs/goal.mli +++ b/proofs/goal.mli @@ -75,6 +75,6 @@ module V82 : sig val nf_evar : Evd.evar_map -> goal -> goal * Evd.evar_map (* Goal represented as a type, doesn't take into account section variables *) - val abstract_type : Evd.evar_map -> goal -> Term.types + val abstract_type : Evd.evar_map -> goal -> EConstr.types end -- cgit v1.2.3 From 02dd160233adc784eac732d97a88356d1f0eaf9b Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 25 Nov 2016 18:34:53 +0100 Subject: Removing various compatibility layers of tactics. --- engine/evd.ml | 7 --- engine/evd.mli | 4 -- interp/constrintern.ml | 9 ++-- interp/constrintern.mli | 20 ++++---- ltac/evar_tactics.ml | 33 +++++++++++-- ltac/evar_tactics.mli | 2 + ltac/extraargs.ml4 | 2 +- ltac/extratactics.ml4 | 22 +-------- ltac/rewrite.ml | 11 +---- ltac/tacinterp.ml | 2 - plugins/cc/ccalgo.ml | 8 +-- plugins/cc/cctac.ml | 10 ++-- plugins/decl_mode/decl_proof_instr.ml | 40 +++++++++------ plugins/derive/derive.ml | 1 + plugins/firstorder/instances.ml | 2 +- plugins/firstorder/rules.ml | 4 +- plugins/firstorder/sequent.ml | 2 + plugins/funind/functional_principles_proofs.ml | 30 +++++------- plugins/funind/functional_principles_types.ml | 4 +- plugins/funind/glob_term_to_relation.ml | 11 +++-- plugins/funind/indfun.ml | 4 +- plugins/funind/invfun.ml | 15 ++---- plugins/funind/recdef.ml | 21 ++++---- plugins/micromega/coq_micromega.ml | 7 ++- plugins/omega/coq_omega.ml | 12 +++-- plugins/quote/quote.ml | 6 +-- plugins/romega/refl_omega.ml | 2 + plugins/rtauto/refl_tauto.ml | 3 +- plugins/setoid_ring/newring.ml | 9 ++-- plugins/ssrmatching/ssrmatching.ml4 | 3 ++ pretyping/cases.ml | 2 +- pretyping/cbv.ml | 2 +- pretyping/cbv.mli | 6 ++- pretyping/coercion.ml | 8 +-- pretyping/pretyping.ml | 8 ++- pretyping/pretyping.mli | 6 +-- pretyping/tacred.ml | 2 +- pretyping/typing.ml | 6 +-- pretyping/typing.mli | 27 ++++++----- pretyping/unification.ml | 2 +- proofs/clenv.ml | 2 +- proofs/evar_refiner.ml | 6 +-- proofs/logic.ml | 2 +- proofs/tacmach.ml | 18 +++---- proofs/tacmach.mli | 65 +++++++++++++------------ stm/lemmas.ml | 1 + tactics/autorewrite.ml | 11 +++-- tactics/class_tactics.ml | 2 - tactics/contradiction.ml | 1 - tactics/eauto.ml | 5 +- tactics/elim.ml | 2 - tactics/eqdecide.ml | 2 - tactics/equality.ml | 20 +++----- tactics/hints.ml | 3 +- tactics/hipattern.ml | 2 +- tactics/inv.ml | 4 +- tactics/leminv.ml | 1 - tactics/tacticals.ml | 22 ++++----- tactics/tactics.ml | 34 +++---------- toplevel/auto_ind_decl.ml | 67 +++++++++++++------------- toplevel/class.ml | 2 +- toplevel/classes.ml | 4 +- toplevel/command.ml | 27 ++++++++--- toplevel/record.ml | 6 ++- toplevel/vernacentries.ml | 3 +- 65 files changed, 343 insertions(+), 344 deletions(-) diff --git a/engine/evd.ml b/engine/evd.ml index 9c05c6c02..6380e3fc1 100644 --- a/engine/evd.ml +++ b/engine/evd.ml @@ -729,13 +729,6 @@ let loc_of_conv_pb evd (pbty,env,t1,t2) = (* excluding defined evars *) -let evar_list c = - let rec evrec acc c = - match kind_of_term c with - | Evar (evk, _ as ev) -> ev :: acc - | _ -> Term.fold_constr evrec acc c in - evrec [] c - let evars_of_term c = let rec evrec acc c = match kind_of_term c with diff --git a/engine/evd.mli b/engine/evd.mli index 4e8284675..699b52c2b 100644 --- a/engine/evd.mli +++ b/engine/evd.mli @@ -420,10 +420,6 @@ val loc_of_conv_pb : evar_map -> evar_constraint -> Loc.t 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 *) diff --git a/interp/constrintern.ml b/interp/constrintern.ml index 8d4d87f2a..e5dd6a6ec 100644 --- a/interp/constrintern.ml +++ b/interp/constrintern.ml @@ -1982,7 +1982,7 @@ let interp_type env sigma ?(impls=empty_internalization_env) c = interp_gen IsType env sigma ~impls c let interp_casted_constr env sigma ?(impls=empty_internalization_env) c typ = - interp_gen (OfType (EConstr.of_constr typ)) env sigma ~impls c + interp_gen (OfType typ) env sigma ~impls c (* Not all evars expected to be resolved *) @@ -2001,7 +2001,7 @@ let interp_constr_evars_impls env evdref ?(impls=empty_internalization_env) c = interp_constr_evars_gen_impls env evdref ~impls WithoutTypeConstraint c let interp_casted_constr_evars_impls env evdref ?(impls=empty_internalization_env) c typ = - interp_constr_evars_gen_impls env evdref ~impls (OfType (EConstr.of_constr typ)) c + interp_constr_evars_gen_impls env evdref ~impls (OfType typ) c let interp_type_evars_impls env evdref ?(impls=empty_internalization_env) c = interp_constr_evars_gen_impls env evdref ~impls IsType c @@ -2093,6 +2093,7 @@ let interp_rawcontext_evars env evdref k bl = let t = understand_tcc_evars env evdref ~expected_type:IsType t' in match b with None -> + let t = EConstr.Unsafe.to_constr t in let d = LocalAssum (na,t) in let impls = if k == Implicit then @@ -2102,7 +2103,9 @@ let interp_rawcontext_evars env evdref k bl = in (push_rel d env, d::params, succ n, impls) | Some b -> - let c = understand_tcc_evars env evdref ~expected_type:(OfType (EConstr.of_constr t)) b in + let c = understand_tcc_evars env evdref ~expected_type:(OfType t) b in + let t = EConstr.Unsafe.to_constr t in + let c = EConstr.Unsafe.to_constr c in let d = LocalDef (na, c, t) in (push_rel d env, d::params, n, impls)) (env,[],k+1,[]) (List.rev bl) diff --git a/interp/constrintern.mli b/interp/constrintern.mli index 61e7c6f5c..2f6795ed4 100644 --- a/interp/constrintern.mli +++ b/interp/constrintern.mli @@ -101,7 +101,7 @@ val interp_constr : env -> evar_map -> ?impls:internalization_env -> constr_expr -> constr Evd.in_evar_universe_context val interp_casted_constr : env -> evar_map -> ?impls:internalization_env -> - constr_expr -> types -> constr Evd.in_evar_universe_context + constr_expr -> EConstr.types -> constr Evd.in_evar_universe_context val interp_type : env -> evar_map -> ?impls:internalization_env -> constr_expr -> types Evd.in_evar_universe_context @@ -109,32 +109,32 @@ val interp_type : env -> evar_map -> ?impls:internalization_env -> (** Main interpretation function expecting all postponed problems to be resolved, but possibly leaving evars. *) -val interp_open_constr : env -> evar_map -> constr_expr -> evar_map * constr +val interp_open_constr : env -> evar_map -> constr_expr -> evar_map * EConstr.constr (** Accepting unresolved evars *) val interp_constr_evars : env -> evar_map ref -> - ?impls:internalization_env -> constr_expr -> constr + ?impls:internalization_env -> constr_expr -> EConstr.constr val interp_casted_constr_evars : env -> evar_map ref -> - ?impls:internalization_env -> constr_expr -> types -> constr + ?impls:internalization_env -> constr_expr -> types -> EConstr.constr val interp_type_evars : env -> evar_map ref -> - ?impls:internalization_env -> constr_expr -> types + ?impls:internalization_env -> constr_expr -> EConstr.types (** Accepting unresolved evars and giving back the manual implicit arguments *) val interp_constr_evars_impls : env -> evar_map ref -> ?impls:internalization_env -> constr_expr -> - constr * Impargs.manual_implicits + EConstr.constr * Impargs.manual_implicits val interp_casted_constr_evars_impls : env -> evar_map ref -> - ?impls:internalization_env -> constr_expr -> types -> - constr * Impargs.manual_implicits + ?impls:internalization_env -> constr_expr -> EConstr.types -> + EConstr.constr * Impargs.manual_implicits val interp_type_evars_impls : env -> evar_map ref -> ?impls:internalization_env -> constr_expr -> - types * Impargs.manual_implicits + EConstr.types * Impargs.manual_implicits (** Interprets constr patterns *) @@ -153,7 +153,7 @@ val interp_reference : ltac_sign -> reference -> glob_constr val interp_binder : env -> evar_map -> Name.t -> constr_expr -> types Evd.in_evar_universe_context -val interp_binder_evars : env -> evar_map ref -> Name.t -> constr_expr -> types +val interp_binder_evars : env -> evar_map ref -> Name.t -> constr_expr -> EConstr.types (** Interpret contexts: returns extended env and context *) diff --git a/ltac/evar_tactics.ml b/ltac/evar_tactics.ml index 6b94da28a..01cff94a8 100644 --- a/ltac/evar_tactics.ml +++ b/ltac/evar_tactics.ml @@ -7,6 +7,9 @@ (************************************************************************) open Util +open Names +open Term +open EConstr open CErrors open Evar_refiner open Tacmach @@ -35,25 +38,32 @@ let instantiate_evar evk (ist,rawc) sigma = let sigma' = w_refine (evk,evi) (lvar ,rawc) sigma in tclEVARS sigma' +let evar_list sigma c = + let rec evrec acc c = + match EConstr.kind sigma c with + | Evar (evk, _ as ev) -> ev :: acc + | _ -> EConstr.fold sigma evrec acc c in + evrec [] c + let instantiate_tac n c ido = Proofview.V82.tactic begin fun gl -> let sigma = gl.sigma in let evl = match ido with - ConclLocation () -> evar_list (pf_concl gl) + ConclLocation () -> evar_list sigma (pf_concl gl) | HypLocation (id,hloc) -> let decl = Environ.lookup_named_val id (Goal.V82.hyps sigma (sig_it gl)) in match hloc with InHyp -> (match decl with - | LocalAssum (_,typ) -> evar_list typ + | LocalAssum (_,typ) -> evar_list sigma (EConstr.of_constr typ) | _ -> error "Please be more specific: in type or value?") | InHypTypeOnly -> - evar_list (NamedDecl.get_type decl) + evar_list sigma (EConstr.of_constr (NamedDecl.get_type decl)) | InHypValueOnly -> (match decl with - | LocalDef (_,body,_) -> evar_list body + | LocalDef (_,body,_) -> evar_list sigma (EConstr.of_constr body) | _ -> error "Not a defined hypothesis.") in if List.length evl < n then error "Not enough uninstantiated existential variables."; @@ -92,3 +102,18 @@ let let_evar name typ = in Sigma (tac, sigma, p) end } + +let hget_evar n = + let open EConstr in + Proofview.Goal.nf_enter { enter = begin fun gl -> + let sigma = Tacmach.New.project gl in + let concl = Proofview.Goal.concl gl in + let evl = evar_list sigma concl in + if List.length evl < n then + error "Not enough uninstantiated existential variables."; + if n <= 0 then error "Incorrect existential variable index."; + let ev = List.nth evl (n-1) in + let ev_type = EConstr.existential_type sigma ev in + Tactics.change_concl (mkLetIn (Anonymous,mkEvar ev,ev_type,concl)) + end } + diff --git a/ltac/evar_tactics.mli b/ltac/evar_tactics.mli index 41266e61c..cfe747665 100644 --- a/ltac/evar_tactics.mli +++ b/ltac/evar_tactics.mli @@ -17,3 +17,5 @@ val instantiate_tac_by_name : Id.t -> Tacinterp.interp_sign * Glob_term.glob_constr -> unit Proofview.tactic val let_evar : Name.t -> EConstr.types -> unit Proofview.tactic + +val hget_evar : int -> unit Proofview.tactic diff --git a/ltac/extraargs.ml4 b/ltac/extraargs.ml4 index f8db0b4fc..53b726432 100644 --- a/ltac/extraargs.ml4 +++ b/ltac/extraargs.ml4 @@ -177,7 +177,7 @@ ARGUMENT EXTEND lglob END let interp_casted_constr ist gl c = - interp_constr_gen (Pretyping.OfType (EConstr.of_constr (pf_concl gl))) ist (pf_env gl) (project gl) c + interp_constr_gen (Pretyping.OfType (pf_concl gl)) ist (pf_env gl) (project gl) c ARGUMENT EXTEND casted_constr TYPED AS constr diff --git a/ltac/extratactics.ml4 b/ltac/extratactics.ml4 index 65c75703b..519633bbe 100644 --- a/ltac/extratactics.ml4 +++ b/ltac/extratactics.ml4 @@ -693,24 +693,8 @@ END hget_evar *) -let hget_evar n = - Proofview.Goal.nf_enter { enter = begin fun gl -> - let sigma = Tacmach.New.project gl in - let concl = Proofview.Goal.concl gl in - let concl = EConstr.Unsafe.to_constr concl in - let evl = evar_list concl in - let concl = EConstr.of_constr concl in - if List.length evl < n then - error "Not enough uninstantiated existential variables."; - if n <= 0 then error "Incorrect existential variable index."; - let ev = List.nth evl (n-1) in - let ev = (fst ev, Array.map EConstr.of_constr (snd ev)) in - let ev_type = existential_type sigma ev in - change_concl (mkLetIn (Anonymous,mkEvar ev,ev_type,concl)) - end } - TACTIC EXTEND hget_evar -| [ "hget_evar" int_or_var(n) ] -> [ hget_evar n ] +| [ "hget_evar" int_or_var(n) ] -> [ Evar_tactics.hget_evar n ] END (**********************************************************************) @@ -747,7 +731,6 @@ let refl_equal = let mkCaseEq a : unit Proofview.tactic = Proofview.Goal.nf_enter { enter = begin fun gl -> let type_of_a = Tacmach.New.of_old (fun g -> Tacmach.pf_unsafe_type_of g a) gl in - let type_of_a = EConstr.of_constr type_of_a in Tacticals.New.tclTHENLIST [Tactics.generalize [(mkApp(EConstr.of_constr (delayed_force refl_equal), [| type_of_a; a|]))]; Proofview.Goal.nf_enter { enter = begin fun gl -> @@ -804,7 +787,6 @@ let destauto t = let destauto_in id = Proofview.Goal.nf_enter { enter = begin fun gl -> let ctype = Tacmach.New.of_old (fun g -> Tacmach.pf_unsafe_type_of g (mkVar id)) gl in - let ctype = EConstr.of_constr ctype in (* Pp.msgnl (Printer.pr_lconstr (mkVar id)); *) (* Pp.msgnl (Printer.pr_lconstr (ctype)); *) destauto ctype @@ -1088,7 +1070,7 @@ END VERNAC COMMAND EXTEND Declare_keys CLASSIFIED AS SIDEFF | [ "Declare" "Equivalent" "Keys" constr(c) constr(c') ] -> [ - let it c = snd (Constrintern.interp_open_constr (Global.env ()) Evd.empty c) in + let it c = EConstr.Unsafe.to_constr (snd (Constrintern.interp_open_constr (Global.env ()) Evd.empty c)) in let k1 = Keys.constr_key (it c) in let k2 = Keys.constr_key (it c') in match k1, k2 with diff --git a/ltac/rewrite.ml b/ltac/rewrite.ml index 3c6bd4563..715929c56 100644 --- a/ltac/rewrite.ml +++ b/ltac/rewrite.ml @@ -135,7 +135,6 @@ let app_poly_check env evars f args = let (evars, cstrs), fc = f evars in let evdref = ref evars in let t = Typing.e_solve_evars env evdref (mkApp (fc, args)) in - let t = EConstr.of_constr t in (!evdref, cstrs), t let app_poly_nocheck env evars f args = @@ -509,7 +508,6 @@ let rec decompose_app_rel env evd t = | App (f, [|arg|]) -> let (f', argl, argr) = decompose_app_rel env evd arg in let ty = Typing.unsafe_type_of env evd argl in - let ty = EConstr.of_constr ty in let f'' = mkLambda (Name default_dependent_ident, ty, mkLambda (Name (Id.of_string "y"), lift 1 ty, mkApp (lift 2 f, [| mkApp (lift 2 f', [| mkRel 2; mkRel 1 |]) |]))) @@ -813,7 +811,6 @@ let resolve_morphism env avoid oldt m ?(fnewt=fun x -> x) args args' (b,cstr) ev let morphargs', morphobjs' = Array.chop first args' in let appm = mkApp(m, morphargs) in let appmtype = Typing.unsafe_type_of env (goalevars evars) appm in - let appmtype = EConstr.of_constr appmtype in let cstrs = List.map (Option.map (fun r -> r.rew_car, get_opt_rew_rel r.rew_prf)) (Array.to_list morphobjs') @@ -1445,7 +1442,6 @@ module Strategies = fun { state ; env ; term1 = t ; ty1 = ty ; cstr ; evars } -> (* let sigma, (c,_) = Tacinterp.interp_open_constr_with_bindings is env (goalevars evars) c in *) let sigma, c = Pretyping.understand_tcc env (goalevars evars) c in - let c = EConstr.of_constr c in let unfolded = try Tacred.try_red_product env sigma c with e when CErrors.noncritical e -> @@ -1693,7 +1689,6 @@ let cl_rewrite_clause_strat strat clause = let apply_glob_constr c l2r occs = (); fun ({ state = () ; env = env } as input) -> let c sigma = let (sigma, c) = Pretyping.understand_tcc env sigma c in - let c = EConstr.of_constr c in (sigma, (c, NoBindings)) in let flags = general_rewrite_unif_flags () in @@ -1702,7 +1697,6 @@ let apply_glob_constr c l2r occs = (); fun ({ state = () ; env = env } as input) let interp_glob_constr_list env = let make c = (); fun sigma -> let sigma, c = Pretyping.understand_tcc env sigma c in - let c = EConstr.of_constr c in (sigma, (c, NoBindings)) in List.map (fun c -> make c, true, None) @@ -1940,7 +1934,6 @@ let build_morphism_signature env sigma m = let m = EConstr.of_constr m in let sigma = Evd.from_ctx ctx in let t = Typing.unsafe_type_of env sigma m in - let t = EConstr.of_constr t in let cstrs = let rec aux t = match EConstr.kind sigma t with @@ -1971,7 +1964,6 @@ let default_morphism sign m = let env = Global.env () in let sigma = Evd.from_env env in let t = Typing.unsafe_type_of env sigma m in - let t = EConstr.of_constr t in let evars, _, sign, cstrs = PropGlobal.build_signature (sigma, Evar.Set.empty) env t (fst sign) (snd sign) in @@ -2111,7 +2103,7 @@ let get_hyp gl (c,l) clause l2r = let env = Tacmach.New.pf_env gl in let sigma, hi = decompose_applied_relation env evars (c,l) in let but = match clause with - | Some id -> EConstr.of_constr (Tacmach.New.pf_get_hyp_typ id gl) + | Some id -> Tacmach.New.pf_get_hyp_typ id gl | None -> nf_evar evars (Tacmach.New.pf_concl gl) in unification_rewrite l2r hi.c1 hi.c2 sigma hi.prf hi.car hi.rel but env @@ -2228,7 +2220,6 @@ let setoid_symmetry_in id = Proofview.V82.tactic (fun gl -> let sigma = project gl in let ctype = pf_unsafe_type_of gl (mkVar id) in - let ctype = EConstr.of_constr ctype in let binders,concl = decompose_prod_assum sigma ctype in let (equiv, args) = decompose_app sigma concl in let rec split_last_two = function diff --git a/ltac/tacinterp.ml b/ltac/tacinterp.ml index d8df07733..36a0336bc 100644 --- a/ltac/tacinterp.ml +++ b/ltac/tacinterp.ml @@ -636,7 +636,6 @@ let interp_gen kind ist allow_patvar flags env sigma (c,ce) = let (evd,c) = catch_error trace (understand_ltac flags env sigma vars kind) c in - let c = EConstr.of_constr c in (* spiwack: to avoid unnecessary modifications of tacinterp, as this function already use effect, I call [run] hoping it doesn't mess up with any assumption. *) @@ -795,7 +794,6 @@ let interp_may_eval f ist env sigma = function let ic = EConstr.Unsafe.to_constr ic in let c = subst_meta [Constr_matching.special_meta,ic] ctxt in let c = Typing.e_solve_evars env evdref (EConstr.of_constr c) in - let c = EConstr.of_constr c in !evdref , c with | Not_found -> diff --git a/plugins/cc/ccalgo.ml b/plugins/cc/ccalgo.ml index 0a980c03b..aedcb7575 100644 --- a/plugins/cc/ccalgo.ml +++ b/plugins/cc/ccalgo.ml @@ -452,8 +452,9 @@ and applist_projection c l = applistc (mkProj (p, hd)) tl) | _ -> applistc c l -let rec canonize_name c = - let func = canonize_name in +let rec canonize_name sigma c = + let c = EConstr.Unsafe.to_constr c in + let func c = canonize_name sigma (EConstr.of_constr c) in match kind_of_term c with | Const (kn,u) -> let canon_const = constant_of_kn (canonical_con kn) in @@ -509,7 +510,7 @@ let rec add_term state t= let b=next uf in let trm = constr_of_term t in let typ = pf_unsafe_type_of state.gls (EConstr.of_constr trm) in - let typ = canonize_name typ in + let typ = canonize_name (project state.gls) typ in let new_node= match t with Symb _ | Product (_,_) -> @@ -833,6 +834,7 @@ let complete_one_class state i= app (Appli(t,Eps id)) (substl [mkVar id] rest) (n-1) in let _c = pf_unsafe_type_of state.gls (EConstr.of_constr (constr_of_term (term state.uf pac.cnode))) in + let _c = EConstr.Unsafe.to_constr _c in let _args = List.map (fun i -> constr_of_term (term state.uf i)) pac.args in diff --git a/plugins/cc/cctac.ml b/plugins/cc/cctac.ml index 62892973d..2ab4dced4 100644 --- a/plugins/cc/cctac.ml +++ b/plugins/cc/cctac.ml @@ -219,7 +219,7 @@ let make_prb gls depth additionnal_terms = | `Nrule patts -> add_quant state id false patts end) (Environ.named_context_of_val (Goal.V82.nf_hyps gls.sigma gls.it)); begin - match atom_of_constr env sigma (EConstr.of_constr (pf_concl gls)) with + match atom_of_constr env sigma (pf_concl gls) with `Eq (t,a,b) -> add_disequality state Goal a b | `Other g -> List.iter @@ -271,7 +271,7 @@ let constr_of_term c = EConstr.of_constr (constr_of_term c) let rec proof_tac p : unit Proofview.tactic = Proofview.Goal.nf_enter { enter = begin fun gl -> - let type_of t = EConstr.of_constr (Tacmach.New.pf_unsafe_type_of gl t) in + let type_of t = Tacmach.New.pf_unsafe_type_of gl t in try (* type_of can raise exceptions *) match p.p_rule with Ax c -> exact_check (EConstr.of_constr c) @@ -343,7 +343,7 @@ let refute_tac c t1 t2 p = let neweq= new_app_global _eq [|intype;tt1;tt2|] in Tacticals.New.tclTHENS (neweq (assert_before (Name hid))) [proof_tac p; simplest_elim false_t] - in refresh_universes (EConstr.of_constr (Tacmach.New.pf_unsafe_type_of gl tt1)) k + in refresh_universes (Tacmach.New.pf_unsafe_type_of gl tt1) k end } let refine_exact_check c gl = @@ -361,7 +361,7 @@ let convert_to_goal_tac c t1 t2 p = let endt=app_global _eq_rect [|sort;tt1;identity;c;tt2;mkVar e|] in Tacticals.New.tclTHENS (neweq (assert_before (Name e))) [proof_tac p; Proofview.V82.tactic (endt refine_exact_check)] - in refresh_universes (EConstr.of_constr (Tacmach.New.pf_unsafe_type_of gl tt2)) k + in refresh_universes (Tacmach.New.pf_unsafe_type_of gl tt2) k end } let convert_to_hyp_tac c1 t1 c2 t2 p = @@ -385,7 +385,7 @@ let discriminate_tac (cstr,u as cstru) p = let trivial = Universes.constr_of_global (Lazy.force _True) in let trivial = EConstr.of_constr trivial in let evm = Tacmach.New.project gl in - let evm, intype = refresh_type env evm (EConstr.of_constr (Tacmach.New.pf_unsafe_type_of gl t1)) in + let evm, intype = refresh_type env evm (Tacmach.New.pf_unsafe_type_of gl t1) in let evm, outtype = Evd.new_sort_variable Evd.univ_flexible evm in let outtype = mkSort outtype in let pred = mkLambda(Name xid,outtype,mkRel 1) in diff --git a/plugins/decl_mode/decl_proof_instr.ml b/plugins/decl_mode/decl_proof_instr.ml index 7123ebcaf..6a0ec3968 100644 --- a/plugins/decl_mode/decl_proof_instr.ml +++ b/plugins/decl_mode/decl_proof_instr.ml @@ -74,7 +74,7 @@ let tcl_change_info_gen info_gen = let concl = pf_concl gls in let hyps = Goal.V82.hyps (project gls) it in let extra = Goal.V82.extra (project gls) it in - let (gl,ev,sigma) = Goal.V82.mk_goal (project gls) hyps (EConstr.of_constr concl) (info_gen extra) in + let (gl,ev,sigma) = Goal.V82.mk_goal (project gls) hyps concl (info_gen extra) in let sigma = Goal.V82.partial_solution sigma it ev in { it = [gl] ; sigma= sigma; } ) @@ -88,7 +88,7 @@ let tcl_erase_info gls = let special_whd gl= let infos=CClosure.create_clos_infos CClosure.all (pf_env gl) in - (fun t -> CClosure.whd_val infos (CClosure.inject t)) + (fun t -> CClosure.whd_val infos (CClosure.inject (EConstr.Unsafe.to_constr t))) let special_nf gl= let infos=CClosure.create_clos_infos CClosure.betaiotazeta (pf_env gl) in @@ -342,7 +342,7 @@ let rec replace_in_list m l = function | c::q -> if Int.equal m (fst c) then l@q else c::replace_in_list m l q let enstack_subsubgoals env se stack gls= - let hd,params = decompose_app (special_whd gls se.se_type) in + let hd,params = decompose_app (special_whd gls (EConstr.of_constr se.se_type)) in match kind_of_term hd with Ind (ind,u as indu) when is_good_inductive env ind -> (* MS: FIXME *) let mib,oib= @@ -397,6 +397,7 @@ let rec nf_list evd = let find_subsubgoal c ctyp skip submetas gls = let env= pf_env gls in let concl = pf_concl gls in + let concl = EConstr.Unsafe.to_constr concl in let evd = mk_evd ((0,concl)::submetas) gls in let stack = Stack.create () in let max_meta = @@ -412,7 +413,7 @@ let find_subsubgoal c ctyp skip submetas gls = try let unifier = Unification.w_unify env se.se_evd Reduction.CUMUL - ~flags:(Unification.elim_flags ()) (EConstr.of_constr ctyp) (EConstr.of_constr se.se_type) in + ~flags:(Unification.elim_flags ()) ctyp (EConstr.of_constr se.se_type) in if n <= 0 then {se with se_evd=meta_assign se.se_meta @@ -433,7 +434,8 @@ let concl_refiner metas body gls = let concl = pf_concl gls in let evd = sig_sig gls in let env = pf_env gls in - let sort = family_of_sort (Typing.e_sort_of env (ref evd) (EConstr.of_constr concl)) in + let sort = family_of_sort (Typing.e_sort_of env (ref evd) concl) in + let concl = EConstr.Unsafe.to_constr concl in let rec aux env avoid subst = function [] -> anomaly ~label:"concl_refiner" (Pp.str "cannot happen") | (n,typ)::rest -> @@ -504,7 +506,7 @@ let mk_stat_or_thesis info gls = function This c -> c | Thesis (For _ ) -> error "\"thesis for ...\" is not applicable here." - | Thesis Plain -> pf_concl gls + | Thesis Plain -> EConstr.Unsafe.to_constr (pf_concl gls) let just_tac _then cut info gls0 = let last_item = @@ -536,7 +538,7 @@ let instr_cut mkstat _thus _then cut gls0 = let c_stat = mkstat info gls0 stat.st_it in let thus_tac gls= if _thus then - thus_tac (mkVar c_id) c_stat [] gls + thus_tac (mkVar c_id) (EConstr.of_constr c_stat) [] gls else tclIDTAC gls in tclTHENS (Proofview.V82.of_tactic (assert_postpone c_id c_stat)) [tclTHEN tcl_erase_info (just_tac _then cut info); @@ -582,7 +584,7 @@ let instr_rew _thus rew_side cut gls0 = | Name id -> id,true in let thus_tac new_eq gls= if _thus then - thus_tac (mkVar c_id) new_eq [] gls + thus_tac (mkVar c_id) (EConstr.of_constr new_eq) [] gls else tclIDTAC gls in match rew_side with Lhs -> @@ -610,7 +612,7 @@ let instr_claim _thus st gls0 = | Name id -> id,true in let thus_tac gls= if _thus then - thus_tac (mkVar id) st.st_it [] gls + thus_tac (mkVar id) (EConstr.of_constr st.st_it) [] gls else tclIDTAC gls in let ninfo1 = {pm_stack= (if _thus then Focus_claim else Claim)::info.pm_stack} in @@ -704,7 +706,7 @@ let instr_suffices _then cut gls0 = let c_ctx,c_head = build_applist c_stat metas in let c_term = applist (mkVar c_id,List.map mkMeta metas) in let thus_tac gls= - thus_tac c_term c_head c_ctx gls in + thus_tac c_term (EConstr.of_constr c_head) c_ctx gls in tclTHENS (Proofview.V82.of_tactic (assert_postpone c_id c_stat)) [tclTHENLIST [ assume_tac ctx; @@ -891,8 +893,9 @@ let build_per_info etype casee gls = let concl=pf_concl gls in let env=pf_env gls in let ctyp=pf_unsafe_type_of gls (EConstr.of_constr casee) in - let is_dep = dependent (project gls) (EConstr.of_constr casee) (EConstr.of_constr concl) in + let is_dep = dependent (project gls) (EConstr.of_constr casee) concl in let hd,args = decompose_app (special_whd gls ctyp) in + let ctyp = EConstr.Unsafe.to_constr ctyp in let (ind,u) = try destInd hd @@ -906,9 +909,10 @@ let build_per_info etype casee gls = let params,real_args = List.chop nparams args in let abstract_obj c body = let typ=pf_unsafe_type_of gls (EConstr.of_constr c) in + let typ = EConstr.Unsafe.to_constr typ in lambda_create env (typ,subst_term (project gls) (EConstr.of_constr c) (EConstr.of_constr body)) in let pred= List.fold_right abstract_obj - real_args (lambda_create env (ctyp,subst_term (project gls) (EConstr.of_constr casee) (EConstr.of_constr concl))) in + real_args (lambda_create env (ctyp,subst_term (project gls) (EConstr.of_constr casee) concl)) in is_dep, {per_casee=casee; per_ctype=ctyp; @@ -963,6 +967,7 @@ let register_nodep_subcase id= function let suppose_tac hyps gls0 = let info = get_its_info gls0 in let thesis = pf_concl gls0 in + let thesis = EConstr.Unsafe.to_constr thesis in let id = pf_get_new_id (Id.of_string "subcase_") gls0 in let clause = build_product (project gls0) hyps thesis in let ninfo1 = {pm_stack=Suppose_case::info.pm_stack} in @@ -1132,7 +1137,7 @@ let rec build_product_dep pat_info per_info args body gls = with Not_found -> snd (st_assoc (Name id) pat_info.pat_aliases) in thesis_for obj typ per_info (pf_env gls) - | Plain -> pf_concl gls in + | Plain -> EConstr.Unsafe.to_constr (pf_concl gls) in mkProd (st.st_label,ptyp,lbody) | [] -> body @@ -1225,6 +1230,7 @@ let pop_stacks stacks = let hrec_for fix_id per_info gls obj_id = let obj=mkVar obj_id in let typ=pf_get_hyp_typ gls obj_id in + let typ = EConstr.Unsafe.to_constr typ in let rc,hd1=decompose_prod typ in let cind,all_args=decompose_app typ in let ind,u = destInd cind in assert (eq_ind ind per_info.per_ind); @@ -1269,14 +1275,16 @@ let rec execute_cases fix_name per_info tacnext args objs nhrec tree gls = let env=pf_env gls in let ctyp=pf_unsafe_type_of gls (EConstr.of_constr casee) in let hd,all_args = decompose_app (special_whd gls ctyp) in + let ctyp = EConstr.Unsafe.to_constr ctyp in let ind', u = destInd hd in let _ = assert (eq_ind ind' ind) in (* just in case *) let params,real_args = List.chop nparams all_args in let abstract_obj c body = let typ=pf_unsafe_type_of gls (EConstr.of_constr c) in + let typ = EConstr.Unsafe.to_constr typ in lambda_create env (typ,subst_term (project gls) (EConstr.of_constr c) (EConstr.of_constr body)) in let elim_pred = List.fold_right abstract_obj - real_args (lambda_create env (ctyp,subst_term (project gls) (EConstr.of_constr casee) (EConstr.of_constr concl))) in + real_args (lambda_create env (ctyp,subst_term (project gls) (EConstr.of_constr casee) concl)) in let case_info = Inductiveops.make_case_info env ind RegularStyle in let gen_arities = Inductive.arities_of_constructors (ind,u) spec in let f_ids typ = @@ -1341,13 +1349,13 @@ let understand_my_constr env sigma c concl = | GEvar _ -> GHole (Loc.ghost,Evar_kinds.QuestionMark Evar_kinds.Expand,Misctypes.IntroAnonymous,None) | rc -> map_glob_constr frob rc in - Pretyping.understand_tcc env sigma ~expected_type:(Pretyping.OfType (EConstr.of_constr concl)) (frob rawc) + Pretyping.understand_tcc env sigma ~expected_type:(Pretyping.OfType concl) (frob rawc) let my_refine c gls = let oc = { run = begin fun sigma -> let sigma = Sigma.to_evar_map sigma in let (sigma, c) = understand_my_constr (pf_env gls) sigma c (pf_concl gls) in - Sigma.Unsafe.of_pair (EConstr.of_constr c, sigma) + Sigma.Unsafe.of_pair (c, sigma) end } in Proofview.V82.of_tactic (Tactics.New.refine oc) gls diff --git a/plugins/derive/derive.ml b/plugins/derive/derive.ml index e39d17b52..f23f4ce7d 100644 --- a/plugins/derive/derive.ml +++ b/plugins/derive/derive.ml @@ -37,6 +37,7 @@ let start_deriving f suchthat lemma = let env' = Environ.push_named (LocalDef (f, ef, f_type)) env in let evdref = ref sigma in let suchthat = Constrintern.interp_type_evars env' evdref suchthat in + let suchthat = EConstr.Unsafe.to_constr suchthat in TCons ( env' , !evdref , suchthat , (fun sigma _ -> TNil sigma)))))) in diff --git a/plugins/firstorder/instances.ml b/plugins/firstorder/instances.ml index 24d4346d9..2881b5333 100644 --- a/plugins/firstorder/instances.ml +++ b/plugins/firstorder/instances.ml @@ -107,7 +107,7 @@ let mk_open_instance id idc gl m t= let typ=pf_unsafe_type_of gl (EConstr.of_constr idc) in (* since we know we will get a product, reduction is not too expensive *) - let (nam,_,_)=destProd (EConstr.Unsafe.to_constr (whd_all env evmap (EConstr.of_constr typ))) in + let (nam,_,_)=destProd (EConstr.Unsafe.to_constr (whd_all env evmap typ)) in match nam with Name id -> id | Anonymous -> dummy_bvid in diff --git a/plugins/firstorder/rules.ml b/plugins/firstorder/rules.ml index bed7a727f..38dae0b20 100644 --- a/plugins/firstorder/rules.ml +++ b/plugins/firstorder/rules.ml @@ -38,14 +38,14 @@ let wrap n b continue seq gls= []->anomaly (Pp.str "Not the expected number of hyps") | nd::q-> let id = NamedDecl.get_id nd in - if occur_var env (project gls) id (EConstr.of_constr (pf_concl gls)) || + if occur_var env (project gls) id (pf_concl gls) || List.exists (occur_var_in_decl env (project gls) id) ctx then (aux (i-1) q (nd::ctx)) else add_formula Hyp (VarRef id) (NamedDecl.get_type nd) (aux (i-1) q (nd::ctx)) gls in let seq1=aux n nc [] in let seq2=if b then - add_formula Concl dummy_id (pf_concl gls) seq1 gls else seq1 in + add_formula Concl dummy_id (EConstr.Unsafe.to_constr (pf_concl gls)) seq1 gls else seq1 in continue seq2 gls let basename_of_global=function diff --git a/plugins/firstorder/sequent.ml b/plugins/firstorder/sequent.ml index 91cd102a2..fb0c22c2b 100644 --- a/plugins/firstorder/sequent.ml +++ b/plugins/firstorder/sequent.ml @@ -201,6 +201,7 @@ let extend_with_ref_list l seq gl = let f gr (seq,gl) = let gl, c = pf_eapply Evd.fresh_global gl gr in let typ=(pf_unsafe_type_of gl (EConstr.of_constr c)) in + let typ = EConstr.Unsafe.to_constr typ in (add_formula Hyp gr typ seq gl,gl) in List.fold_right f l (seq,gl) @@ -216,6 +217,7 @@ let extend_with_auto_hints l seq gl= (try let (gr, _) = Termops.global_of_constr (project gl) c in let typ=(pf_unsafe_type_of gl c) in + let typ = EConstr.Unsafe.to_constr typ in seqref:=add_formula Hint gr typ !seqref gl with Not_found->()) | _-> () in diff --git a/plugins/funind/functional_principles_proofs.ml b/plugins/funind/functional_principles_proofs.ml index 656924e38..f4fa61a22 100644 --- a/plugins/funind/functional_principles_proofs.ml +++ b/plugins/funind/functional_principles_proofs.ml @@ -394,7 +394,7 @@ let rewrite_until_var arg_num eq_ids : tactic = *) let test_var g = let sigma = project g in - let _,args = destApp sigma (EConstr.of_constr (pf_concl g)) in + let _,args = destApp sigma (pf_concl g) in not ((isConstruct sigma args.(arg_num)) || isAppConstruct sigma args.(arg_num)) in let rec do_rewrite eq_ids g = @@ -551,7 +551,7 @@ let clean_hyp_with_heq ptes_infos eq_hyps hyp_id env sigma = tclIDTAC in try - scan_type [] (EConstr.of_constr (Typing.unsafe_type_of env sigma (mkVar hyp_id))), [hyp_id] + scan_type [] (Typing.unsafe_type_of env sigma (mkVar hyp_id)), [hyp_id] with TOREMOVE -> thin [hyp_id],[] @@ -602,7 +602,6 @@ let treat_new_case ptes_infos nb_prod continue_tac term dyn_infos = observe_tac "after_introduction" (fun g' -> (* We get infos on the equations introduced*) let new_term_value_eq = pf_unsafe_type_of g' (mkVar heq_id) in - let new_term_value_eq = EConstr.of_constr new_term_value_eq in (* compute the new value of the body *) let new_term_value = match EConstr.kind (project g') new_term_value_eq with @@ -615,7 +614,7 @@ let treat_new_case ptes_infos nb_prod continue_tac term dyn_infos = in let fun_body = mkLambda(Anonymous, - EConstr.of_constr (pf_unsafe_type_of g' term), + pf_unsafe_type_of g' term, Termops.replace_term (project g') term (mkRel 1) dyn_infos.info ) in @@ -708,9 +707,8 @@ let build_proof let t = dyn_info'.info in let dyn_infos = {dyn_info' with info = mkCase(ci,ct,t,cb)} in - let g_nb_prod = nb_prod (project g) (EConstr.of_constr (pf_concl g)) in + let g_nb_prod = nb_prod (project g) (pf_concl g) in let type_of_term = pf_unsafe_type_of g t in - let type_of_term = EConstr.of_constr type_of_term in let term_eq = make_refl_eq (Lazy.force refl_equal) type_of_term t in @@ -722,7 +720,7 @@ let build_proof (fun g -> observe_tac "toto" ( tclTHENSEQ [Proofview.V82.of_tactic (Simple.case t); (fun g' -> - let g'_nb_prod = nb_prod (project g') (EConstr.of_constr (pf_concl g')) in + let g'_nb_prod = nb_prod (project g') (pf_concl g') in let nb_instanciate_partial = g'_nb_prod - g_nb_prod in observe_tac "treat_new_case" (treat_new_case @@ -742,7 +740,7 @@ let build_proof build_proof do_finalize_t {dyn_infos with info = t} g | Lambda(n,t,b) -> begin - match EConstr.kind sigma (EConstr.of_constr ( pf_concl g)) with + match EConstr.kind sigma (pf_concl g) with | Prod _ -> tclTHEN (Proofview.V82.of_tactic intro) @@ -914,7 +912,7 @@ let prove_rec_hyp_for_struct fix_info = (fun eq_hyps -> tclTHEN (rewrite_until_var (fix_info.idx) eq_hyps) (fun g -> - let _,pte_args = destApp (project g) (EConstr.of_constr (pf_concl g)) in + let _,pte_args = destApp (project g) (pf_concl g) in let rec_hyp_proof = mkApp(mkVar fix_info.name,array_get_start pte_args) in @@ -938,7 +936,7 @@ let generalize_non_dep hyp g = let hyp = get_id decl in if Id.List.mem hyp hyps || List.exists (Termops.occur_var_in_decl env (project g) hyp) keep - || Termops.occur_var env (project g) hyp (EConstr.of_constr hyp_typ) + || Termops.occur_var env (project g) hyp hyp_typ || Termops.is_section_variable hyp (* should be dangerous *) then (clear,decl::keep) else (hyp::clear,keep)) @@ -1054,7 +1052,7 @@ let do_replace (evd:Evd.evar_map ref) params rec_arg_num rev_args_id f fun_num a let _ = Typing.e_type_of ~refresh:true (Global.env ()) evd res in res in - let nb_intro_to_do = nb_prod (project g) (EConstr.of_constr (pf_concl g)) in + let nb_intro_to_do = nb_prod (project g) (pf_concl g) in tclTHEN (tclDO nb_intro_to_do (Proofview.V82.of_tactic intro)) ( @@ -1070,7 +1068,6 @@ let do_replace (evd:Evd.evar_map ref) params rec_arg_num rev_args_id f fun_num a let prove_princ_for_struct (evd:Evd.evar_map ref) interactive_proof fun_num fnames all_funs _nparams : tactic = fun g -> let princ_type = pf_concl g in - let princ_type = EConstr.of_constr princ_type in (* Pp.msgnl (str "princ_type " ++ Printer.pr_lconstr princ_type); *) (* Pp.msgnl (str "all_funs "); *) (* Array.iter (fun c -> Pp.msgnl (Printer.pr_lconstr c)) all_funs; *) @@ -1258,7 +1255,7 @@ let prove_princ_for_struct (evd:Evd.evar_map ref) interactive_proof fun_num fnam in let intros_after_fixes : tactic = fun gl -> - let ctxt,pte_app = (decompose_prod_assum (project gl) (EConstr.of_constr (pf_concl gl))) in + let ctxt,pte_app = (decompose_prod_assum (project gl) (pf_concl gl)) in let pte,pte_args = (decompose_app (project gl) pte_app) in try let pte = @@ -1431,12 +1428,12 @@ let backtrack_eqs_until_hrec hrec eqs : tactic = let rewrite = tclFIRST (List.map (fun x -> Proofview.V82.of_tactic (Equality.rewriteRL x)) eqs ) in - let _,hrec_concl = decompose_prod (project gls) (EConstr.of_constr (pf_unsafe_type_of gls (mkVar hrec))) in + let _,hrec_concl = decompose_prod (project gls) (pf_unsafe_type_of gls (mkVar hrec)) in let f_app = Array.last (snd (destApp (project gls) hrec_concl)) in let f = (fst (destApp (project gls) f_app)) in let rec backtrack : tactic = fun g -> - let f_app = Array.last (snd (destApp (project g) (EConstr.of_constr (pf_concl g)))) in + let f_app = Array.last (snd (destApp (project g) (pf_concl g))) in match EConstr.kind (project g) f_app with | App(f',_) when eq_constr (project g) f' f -> tclIDTAC g | _ -> tclTHEN rewrite backtrack g @@ -1525,7 +1522,6 @@ let prove_principle_for_gen (f_ref,functional_ref,eq_ref) tcc_lemma_ref is_mes rec_arg_num rec_arg_type relation gl = let princ_type = pf_concl gl in - let princ_type = EConstr.of_constr princ_type in let princ_info = compute_elim_sig (project gl) princ_type in let fresh_id = let avoid = ref (pf_ids_of_hyps gl) in @@ -1664,7 +1660,7 @@ let prove_principle_for_gen Proofview.V82.of_tactic (Equality.rewriteLR (mkConst eq_ref)); (* observe_tac "finish" *) (fun gl' -> let body = - let _,args = destApp (project gl') (EConstr.of_constr (pf_concl gl')) in + let _,args = destApp (project gl') (pf_concl gl') in Array.last args in let body_info rec_hyps = diff --git a/plugins/funind/functional_principles_types.ml b/plugins/funind/functional_principles_types.ml index d964002f9..ba01b3b04 100644 --- a/plugins/funind/functional_principles_types.ml +++ b/plugins/funind/functional_principles_types.ml @@ -493,7 +493,7 @@ let make_scheme evd (fas : (pconstant*glob_sort) list) : Safe_typing.private_con in let _ = evd := sigma in let l_schemes = - List.map (EConstr.of_constr %> Typing.unsafe_type_of env sigma) schemes + List.map (EConstr.of_constr %> Typing.unsafe_type_of env sigma %> EConstr.Unsafe.to_constr) schemes in let i = ref (-1) in let sorts = @@ -671,7 +671,7 @@ let build_case_scheme fa = Indrec.build_case_analysis_scheme_default env sigma ind sf in let sigma = Sigma.to_evar_map sigma in - let scheme_type = (Typing.unsafe_type_of env sigma) (EConstr.of_constr scheme) in + let scheme_type = EConstr.Unsafe.to_constr ((Typing.unsafe_type_of env sigma) (EConstr.of_constr scheme)) in let sorts = (fun (_,_,x) -> Universes.new_sort_in_family (Pretyping.interp_elimination_sort x) diff --git a/plugins/funind/glob_term_to_relation.ml b/plugins/funind/glob_term_to_relation.ml index fc5a287ae..fd2f4bbd3 100644 --- a/plugins/funind/glob_term_to_relation.ml +++ b/plugins/funind/glob_term_to_relation.ml @@ -504,6 +504,7 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return = *) let rt_as_constr,ctx = Pretyping.understand env (Evd.from_env env) rt in let rt_typ = Typing.unsafe_type_of env (Evd.from_env env) (EConstr.of_constr rt_as_constr) in + let rt_typ = EConstr.Unsafe.to_constr rt_typ in let res_raw_type = Detyping.detype false [] env (Evd.from_env env) rt_typ in let res = fresh_id args_res.to_avoid "_res" in let new_avoid = res::args_res.to_avoid in @@ -612,6 +613,7 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return = let v_res = build_entry_lc env funnames avoid v in let v_as_constr,ctx = Pretyping.understand env (Evd.from_env env) v in let v_type = Typing.unsafe_type_of env (Evd.from_env env) (EConstr.of_constr v_as_constr) in + let v_type = EConstr.Unsafe.to_constr v_type in let new_env = match n with Anonymous -> env @@ -629,7 +631,7 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return = let b_as_constr,ctx = Pretyping.understand env (Evd.from_env env) b in let b_typ = Typing.unsafe_type_of env (Evd.from_env env) (EConstr.of_constr b_as_constr) in let (ind,_) = - try Inductiveops.find_inductive env (Evd.from_env env) (EConstr.of_constr b_typ) + try Inductiveops.find_inductive env (Evd.from_env env) b_typ with Not_found -> user_err (str "Cannot find the inductive associated to " ++ Printer.pr_glob_constr b ++ str " in " ++ @@ -661,7 +663,7 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return = let b_as_constr,ctx = Pretyping.understand env (Evd.from_env env) b in let b_typ = Typing.unsafe_type_of env (Evd.from_env env) (EConstr.of_constr b_as_constr) in let (ind,_) = - try Inductiveops.find_inductive env (Evd.from_env env) (EConstr.of_constr b_typ) + try Inductiveops.find_inductive env (Evd.from_env env) b_typ with Not_found -> user_err (str "Cannot find the inductive associated to " ++ Printer.pr_glob_constr b ++ str " in " ++ @@ -706,7 +708,7 @@ and build_entry_lc_from_case env funname make_discr let types = List.map (fun (case_arg,_) -> let case_arg_as_constr,ctx = Pretyping.understand env (Evd.from_env env) case_arg in - Typing.unsafe_type_of env (Evd.from_env env) (EConstr.of_constr case_arg_as_constr) + EConstr.Unsafe.to_constr (Typing.unsafe_type_of env (Evd.from_env env) (EConstr.of_constr case_arg_as_constr)) ) el in (****** The next works only if the match is not dependent ****) @@ -755,6 +757,7 @@ and build_entry_lc_from_case_term env types funname make_discr patterns_to_preve let typ_of_id = Typing.unsafe_type_of env_with_pat_ids (Evd.from_env env) (EConstr.mkVar id) in + let typ_of_id = EConstr.Unsafe.to_constr typ_of_id in let raw_typ_of_id = Detyping.detype false [] env_with_pat_ids (Evd.from_env env) typ_of_id @@ -808,6 +811,7 @@ and build_entry_lc_from_case_term env types funname make_discr patterns_to_preve if Id.Set.mem id this_pat_ids then (Prod (Name id), let typ_of_id = Typing.unsafe_type_of new_env (Evd.from_env env) (EConstr.mkVar id) in + let typ_of_id = EConstr.Unsafe.to_constr typ_of_id in let raw_typ_of_id = Detyping.detype false [] new_env (Evd.from_env env) typ_of_id in @@ -1122,6 +1126,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt = let t',ctx = Pretyping.understand env evd t in let evd = Evd.from_ctx ctx in let type_t' = Typing.unsafe_type_of env evd (EConstr.of_constr t') in + let type_t' = EConstr.Unsafe.to_constr type_t' in let new_env = Environ.push_rel (LocalDef (n,t',type_t')) env in let new_b,id_to_exclude = rebuild_cons new_env diff --git a/plugins/funind/indfun.ml b/plugins/funind/indfun.ml index e22fed391..1cde4420e 100644 --- a/plugins/funind/indfun.ml +++ b/plugins/funind/indfun.ml @@ -78,11 +78,11 @@ let functional_induction with_clean c princl pat = ++Printer.pr_leconstr (mkConst c') ) in let princ = EConstr.of_constr princ in - (princ,NoBindings,EConstr.of_constr (Tacmach.pf_unsafe_type_of g' princ),g') + (princ,NoBindings,Tacmach.pf_unsafe_type_of g' princ,g') | _ -> raise (UserError(None,str "functional induction must be used with a function" )) end | Some ((princ,binding)) -> - princ,binding,EConstr.of_constr (Tacmach.pf_unsafe_type_of g princ),g + princ,binding,Tacmach.pf_unsafe_type_of g princ,g in let sigma = Tacmach.project g' in let princ_infos = Tactics.compute_elim_sig (Tacmach.project g') princ_type in diff --git a/plugins/funind/invfun.ml b/plugins/funind/invfun.ml index be82010d9..5cbec7743 100644 --- a/plugins/funind/invfun.ml +++ b/plugins/funind/invfun.ml @@ -263,7 +263,7 @@ let prove_fun_correct evd functional_induction funs_constr graphs_constr schemes let princ_type = nf_zeta princ_type in let princ_infos = Tactics.compute_elim_sig evd princ_type in (* The number of args of the function is then easily computable *) - let nb_fun_args = nb_prod (project g) (EConstr.of_constr (pf_concl g)) - 2 in + let nb_fun_args = nb_prod (project g) (pf_concl g) - 2 in let args_names = generate_fresh_id (Id.of_string "x") [] nb_fun_args in let ids = args_names@(pf_ids_of_hyps g) in (* Since we cannot ensure that the functional principle is defined in the @@ -315,7 +315,6 @@ let prove_fun_correct evd functional_induction funs_constr graphs_constr schemes List.fold_right (fun hid acc -> let type_of_hid = pf_unsafe_type_of g (mkVar hid) in - let type_of_hid = EConstr.of_constr type_of_hid in let sigma = project g in match EConstr.kind sigma type_of_hid with | Prod(_,_,t') -> @@ -503,7 +502,7 @@ and intros_with_rewrite_aux : tactic = fun g -> let eq_ind = make_eq () in let sigma = project g in - match EConstr.kind sigma (EConstr.of_constr (pf_concl g)) with + match EConstr.kind sigma (pf_concl g) with | Prod(_,t,t') -> begin match EConstr.kind sigma t with @@ -591,7 +590,7 @@ and intros_with_rewrite_aux : tactic = let rec reflexivity_with_destruct_cases g = let destruct_case () = try - match EConstr.kind (project g) (snd (destApp (project g) (EConstr.of_constr (pf_concl g)))).(2) with + match EConstr.kind (project g) (snd (destApp (project g) (pf_concl g))).(2) with | Case(_,_,v,_) -> tclTHENSEQ[ Proofview.V82.of_tactic (simplest_case v); @@ -608,7 +607,7 @@ let rec reflexivity_with_destruct_cases g = match sc with None -> tclIDTAC g | Some id -> - match EConstr.kind (project g) (EConstr.of_constr (pf_unsafe_type_of g (mkVar id))) with + match EConstr.kind (project g) (pf_unsafe_type_of g (mkVar id)) with | App(eq,[|_;t1;t2|]) when EConstr.eq_constr (project g) eq eq_ind -> if Equality.discriminable (pf_env g) (project g) t1 t2 then Proofview.V82.of_tactic (Equality.discrHyp id) g @@ -674,12 +673,11 @@ let prove_fun_complete funcs graphs schemes lemmas_types_infos i : tactic = let f = funcs.(i) in let graph_principle = nf_zeta (EConstr.of_constr schemes.(i)) in let princ_type = pf_unsafe_type_of g graph_principle in - let princ_type = EConstr.of_constr princ_type in let princ_infos = Tactics.compute_elim_sig (project g) princ_type in (* Then we get the number of argument of the function and compute a fresh name for each of them *) - let nb_fun_args = nb_prod (project g) (EConstr.of_constr (pf_concl g)) - 2 in + let nb_fun_args = nb_prod (project g) (pf_concl g) - 2 in let args_names = generate_fresh_id (Id.of_string "x") [] nb_fun_args in let ids = args_names@(pf_ids_of_hyps g) in (* and fresh names for res H and the principle (cf bug bug #1174) *) @@ -937,7 +935,6 @@ let derive_correctness make_scheme functional_induction (funs: pconstant list) ( let revert_graph kn post_tac hid g = let sigma = project g in let typ = pf_unsafe_type_of g (mkVar hid) in - let typ = EConstr.of_constr typ in match EConstr.kind sigma typ with | App(i,args) when isInd sigma i -> let ((kn',num) as ind'),u = destInd sigma i in @@ -990,7 +987,6 @@ let functional_inversion kn hid fconst f_correct : tactic = let old_ids = List.fold_right Id.Set.add (pf_ids_of_hyps g) Id.Set.empty in let sigma = project g in let type_of_h = pf_unsafe_type_of g (mkVar hid) in - let type_of_h = EConstr.of_constr type_of_h in match EConstr.kind sigma type_of_h with | App(eq,args) when EConstr.eq_constr sigma eq (make_eq ()) -> let pre_tac,f_args,res = @@ -1044,7 +1040,6 @@ let invfun qhyp f g = (fun hid -> Proofview.V82.tactic begin fun g -> let sigma = project g in let hyp_typ = pf_unsafe_type_of g (mkVar hid) in - let hyp_typ = EConstr.of_constr hyp_typ in match EConstr.kind sigma hyp_typ with | App(eq,args) when EConstr.eq_constr sigma eq (make_eq ()) -> begin diff --git a/plugins/funind/recdef.ml b/plugins/funind/recdef.ml index a80a7b5e7..adbdb1eb7 100644 --- a/plugins/funind/recdef.ml +++ b/plugins/funind/recdef.ml @@ -123,7 +123,7 @@ let pf_get_new_ids idl g = let compute_renamed_type gls c = EConstr.of_constr (rename_bound_vars_as_displayed (*no avoid*) [] (*no rels*) [] - (pf_unsafe_type_of gls c)) + (EConstr.Unsafe.to_constr (pf_unsafe_type_of gls c))) let h'_id = Id.of_string "h'" let teq_id = Id.of_string "teq" let ano_id = Id.of_string "anonymous" @@ -412,7 +412,6 @@ let treat_case forbid_new_ids to_intros finalize_tac nb_lam e infos : tactic = h_intros to_intros; (fun g' -> let ty_teq = pf_unsafe_type_of g' (mkVar heq) in - let ty_teq = EConstr.of_constr ty_teq in let teq_lhs,teq_rhs = let _,args = try destApp (project g') ty_teq with DestKO -> assert false in args.(1),args.(2) @@ -522,19 +521,19 @@ let rec prove_lt hyple g = let sigma = project g in begin try - let (varx,varz) = match decompose_app sigma (EConstr.of_constr (pf_concl g)) with + let (varx,varz) = match decompose_app sigma (pf_concl g) with | _, x::z::_ when isVar sigma x && isVar sigma z -> x, z | _ -> assert false in let h = List.find (fun id -> - match decompose_app sigma (EConstr.of_constr (pf_unsafe_type_of g (mkVar id))) with + match decompose_app sigma (pf_unsafe_type_of g (mkVar id)) with | _, t::_ -> EConstr.eq_constr sigma t varx | _ -> false ) hyple in let y = - List.hd (List.tl (snd (decompose_app sigma (EConstr.of_constr (pf_unsafe_type_of g (mkVar h)))))) in + List.hd (List.tl (snd (decompose_app sigma (pf_unsafe_type_of g (mkVar h))))) in observe_tclTHENLIST (str "prove_lt1")[ Proofview.V82.of_tactic (apply (mkApp(le_lt_trans (),[|varx;y;varz;mkVar h|]))); observe_tac (str "prove_lt") (prove_lt hyple) @@ -698,7 +697,6 @@ let mkDestructEq : then None else Some id) hyps in let to_revert_constr = List.rev_map mkVar to_revert in let type_of_expr = pf_unsafe_type_of g expr in - let type_of_expr = EConstr.of_constr type_of_expr in let new_hyps = mkApp(Lazy.force refl_equal, [|type_of_expr; expr|]):: to_revert_constr in pf_typel new_hyps (fun _ -> @@ -707,7 +705,7 @@ let mkDestructEq : (fun g2 -> let changefun patvars = { run = fun sigma -> let redfun = pattern_occs [Locus.AllOccurrencesBut [1], expr] in - let Sigma (c, sigma, p) = redfun.Reductionops.e_redfun (pf_env g2) sigma (EConstr.of_constr (pf_concl g2)) in + let Sigma (c, sigma, p) = redfun.Reductionops.e_redfun (pf_env g2) sigma (pf_concl g2) in Sigma (c, sigma, p) } in Proofview.V82.of_tactic (change_in_concl None changefun) g2); @@ -846,7 +844,7 @@ let equation_case next_step (ci,a,t,l) expr_info continuation_tac infos = let rec prove_le g = let sigma = project g in let x,z = - let _,args = decompose_app sigma (EConstr.of_constr (pf_concl g)) in + let _,args = decompose_app sigma (pf_concl g) in (List.hd args,List.hd (List.tl args)) in tclFIRST[ @@ -857,9 +855,8 @@ let rec prove_le g = let matching_fun = pf_is_matching g (Pattern.PApp(Pattern.PRef (reference_of_constr (EConstr.Unsafe.to_constr (le ()))),[|Pattern.PVar (destVar sigma x);Pattern.PMeta None|])) in - let (h,t) = List.find (fun (_,t) -> matching_fun (EConstr.of_constr t)) (pf_hyps_types g) + let (h,t) = List.find (fun (_,t) -> matching_fun t) (pf_hyps_types g) in - let t = EConstr.of_constr t in let y = let _,args = decompose_app sigma t in List.hd (List.tl args) @@ -1350,7 +1347,7 @@ let open_new_goal build_proof sigma using_lemmas ref_ goal_name (gls_type,decomp ] gls) (fun g -> let sigma = project g in - match EConstr.kind sigma (EConstr.of_constr (pf_concl g)) with + match EConstr.kind sigma (pf_concl g) with | App(f,_) when EConstr.eq_constr sigma f (well_founded ()) -> Proofview.V82.of_tactic (Auto.h_auto None [] (Some [])) g | _ -> @@ -1523,9 +1520,11 @@ let recursive_definition is_mes function_name rec_impls type_of_f r rec_arg_num let env = Global.env() in let evd = ref (Evd.from_env env) in let function_type = interp_type_evars env evd type_of_f in + let function_type = EConstr.Unsafe.to_constr function_type in let env = push_named (Context.Named.Declaration.LocalAssum (function_name,function_type)) env in (* Pp.msgnl (str "function type := " ++ Printer.pr_lconstr function_type); *) let ty = interp_type_evars env evd ~impls:rec_impls eq in + let ty = EConstr.Unsafe.to_constr ty in let evm, nf = Evarutil.nf_evars_and_universes !evd in let equation_lemma_type = nf_betaiotazeta (EConstr.of_constr (nf ty)) in let function_type = nf function_type in diff --git a/plugins/micromega/coq_micromega.ml b/plugins/micromega/coq_micromega.ml index ced572466..a2346cc90 100644 --- a/plugins/micromega/coq_micromega.ml +++ b/plugins/micromega/coq_micromega.ml @@ -1606,6 +1606,7 @@ let rec parse_hyps gl parse_arith env tg hyps = match hyps with | [] -> ([],env,tg) | (i,t)::l -> + let t = EConstr.Unsafe.to_constr t in let (lhyps,env,tg) = parse_hyps gl parse_arith env tg l in try let (c,env,tg) = parse_formula gl parse_arith env tg t in @@ -1713,7 +1714,7 @@ let micromega_order_change spec cert cert_typ env ff (*: unit Proofview.tactic* ("__varmap", vm, Term.mkApp( coq_VarMap, [|spec.typ|])); ("__wit", cert, cert_typ) ] - (Tacmach.pf_concl gl)) + (EConstr.Unsafe.to_constr (Tacmach.pf_concl gl))) ] end } @@ -1964,6 +1965,7 @@ let micromega_gen Proofview.Goal.nf_enter { enter = begin fun gl -> let gl = Tacmach.New.of_old (fun x -> x) gl in let concl = Tacmach.pf_concl gl in + let concl = EConstr.Unsafe.to_constr concl in let hyps = Tacmach.pf_hyps_types gl in try let (hyps,concl,env) = parse_goal gl parse_arith Env.empty hyps concl in @@ -2051,7 +2053,7 @@ let micromega_order_changer cert env ff = [["Coq" ; "micromega" ; "VarMap"] ; ["VarMap"]] "t", [|typ|])); ("__wit", cert, cert_typ) ] - (Tacmach.pf_concl gl))); + (EConstr.Unsafe.to_constr (Tacmach.pf_concl gl)))); (* Tacticals.New.tclTHENLIST (List.map (fun id -> (Tactics.introduction id)) ids)*) ] end } @@ -2072,6 +2074,7 @@ let micromega_genr prover tac = Proofview.Goal.nf_enter { enter = begin fun gl -> let gl = Tacmach.New.of_old (fun x -> x) gl in let concl = Tacmach.pf_concl gl in + let concl = EConstr.Unsafe.to_constr concl in let hyps = Tacmach.pf_hyps_types gl in try diff --git a/plugins/omega/coq_omega.ml b/plugins/omega/coq_omega.ml index 51790f4c9..665486272 100644 --- a/plugins/omega/coq_omega.ml +++ b/plugins/omega/coq_omega.ml @@ -38,9 +38,9 @@ open OmegaSolver let elim_id id = Proofview.Goal.nf_enter { enter = begin fun gl -> - simplest_elim (EConstr.of_constr (Tacmach.New.pf_global id gl)) + simplest_elim (Tacmach.New.pf_global id gl) end } -let resolve_id id gl = Proofview.V82.of_tactic (apply (EConstr.of_constr (pf_global gl id))) gl +let resolve_id id gl = Proofview.V82.of_tactic (apply (pf_global gl id)) gl let timing timer_name f arg = f arg @@ -568,7 +568,7 @@ let abstract_path typ path t = mkLambda (Name (Id.of_string "x"), typ, abstract), !term_occur let focused_simpl path gl = - let newc = context (fun i t -> EConstr.Unsafe.to_constr (pf_nf gl (EConstr.of_constr t))) (List.rev path) (pf_concl gl) in + let newc = context (fun i t -> EConstr.Unsafe.to_constr (pf_nf gl (EConstr.of_constr t))) (List.rev path) (EConstr.Unsafe.to_constr (pf_concl gl)) in let newc = EConstr.of_constr newc in Proofview.V82.of_tactic (convert_concl_no_check newc DEFAULTcast) gl @@ -632,6 +632,7 @@ let mkNewMeta () = mkMeta (Evarutil.new_meta()) let clever_rewrite_base_poly typ p result theorem gl = let full = pf_concl gl in + let full = EConstr.Unsafe.to_constr full in let (abstracted,occ) = abstract_path typ (List.rev p) full in let t = applist @@ -663,6 +664,7 @@ let clever_rewrite_gen_nat p result (t,args) = let clever_rewrite p vpath t gl = let full = pf_concl gl in + let full = EConstr.Unsafe.to_constr full in let (abstracted,occ) = abstract_path (Lazy.force coq_Z) (List.rev p) full in let vargs = List.map (fun p -> occurrence p occ) vpath in let t' = applist(t, (vargs @ [abstracted])) in @@ -1424,6 +1426,7 @@ let coq_omega = Proofview.Goal.nf_enter { enter = begin fun gl -> clear_constr_tables (); let hyps_types = Tacmach.New.pf_hyps_types gl in + let hyps_types = List.map (on_snd EConstr.Unsafe.to_constr) hyps_types in let destructure_omega = Tacmach.New.of_old destructure_omega gl in let tactic_normalisation, system = List.fold_left destructure_omega ([],[]) hyps_types in @@ -1607,6 +1610,7 @@ let nat_inject = with e when catchable_exception e -> loop lit end in let hyps_types = Tacmach.New.pf_hyps_types gl in + let hyps_types = List.map (on_snd EConstr.Unsafe.to_constr) hyps_types in loop (List.rev hyps_types) end } @@ -1722,7 +1726,7 @@ let destructure_hyps = | Kimp(t1,t2) -> (* t1 and t2 might be in Type rather than Prop. For t1, the decidability check will ensure being Prop. *) - if is_Prop (type_of (EConstr.of_constr t2)) + if is_Prop (EConstr.Unsafe.to_constr (type_of (EConstr.of_constr t2))) then let d1 = decidability t1 in Tacticals.New.tclTHENLIST [ diff --git a/plugins/quote/quote.ml b/plugins/quote/quote.ml index 87276f5df..edf34607b 100644 --- a/plugins/quote/quote.ml +++ b/plugins/quote/quote.ml @@ -459,8 +459,7 @@ let quote f lid = let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in let f = Tacmach.New.pf_global f gl in - let f = EConstr.of_constr f in - let cl = List.map (fun id -> Tacmach.New.pf_global id gl) lid in + let cl = List.map (fun id -> EConstr.to_constr sigma (Tacmach.New.pf_global id gl)) lid in let ivs = compute_ivs f cl gl in let concl = Proofview.Goal.concl gl in let quoted_terms = quote_terms env sigma ivs [concl] in @@ -478,8 +477,7 @@ let gen_quote cont c f lid = let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in let f = Tacmach.New.pf_global f gl in - let f = EConstr.of_constr f in - let cl = List.map (fun id -> Tacmach.New.pf_global id gl) lid in + let cl = List.map (fun id -> EConstr.to_constr sigma (Tacmach.New.pf_global id gl)) lid in let ivs = compute_ivs f cl gl in let quoted_terms = quote_terms env sigma ivs [c] in let (p, vm) = match quoted_terms with diff --git a/plugins/romega/refl_omega.ml b/plugins/romega/refl_omega.ml index ab5033601..cfe14b230 100644 --- a/plugins/romega/refl_omega.ml +++ b/plugins/romega/refl_omega.ml @@ -740,6 +740,7 @@ and oproposition_of_constr env ((negated,depends,origin,path) as ctxt) gl c = let reify_gl env gl = let concl = Tacmach.pf_concl gl in + let concl = EConstr.Unsafe.to_constr concl in let t_concl = Pnot (oproposition_of_constr env (true,[],id_concl,[O_mono]) gl concl) in if !debug then begin @@ -748,6 +749,7 @@ let reify_gl env gl = end; let rec loop = function (i,t) :: lhyps -> + let t = EConstr.Unsafe.to_constr t in let t' = oproposition_of_constr env (false,[],i,[]) gl t in if !debug then begin Printf.printf " %s: " (Names.Id.to_string i); diff --git a/plugins/rtauto/refl_tauto.ml b/plugins/rtauto/refl_tauto.ml index 981ce2a61..475380512 100644 --- a/plugins/rtauto/refl_tauto.ml +++ b/plugins/rtauto/refl_tauto.ml @@ -262,8 +262,9 @@ let rtauto_tac gls= let gl=pf_concl gls in let _= if Retyping.get_sort_family_of - (pf_env gls) (Tacmach.project gls) (EConstr.of_constr gl) != InProp + (pf_env gls) (Tacmach.project gls) gl != InProp then user_err ~hdr:"rtauto" (Pp.str "goal should be in Prop") in + let gl = EConstr.Unsafe.to_constr gl in let glf=make_form gamma gls gl in let hyps=make_hyps gamma gls [gl] (pf_hyps gls) in let formula= diff --git a/plugins/setoid_ring/newring.ml b/plugins/setoid_ring/newring.ml index 358ea5685..f52557095 100644 --- a/plugins/setoid_ring/newring.ml +++ b/plugins/setoid_ring/newring.ml @@ -142,7 +142,7 @@ let ic c = let env = Global.env() in let sigma = Evd.from_env env in let sigma, c = Constrintern.interp_open_constr env sigma c in - (sigma, EConstr.of_constr c) + (sigma, c) let ic_unsafe c = (*FIXME remove *) let env = Global.env() in @@ -505,8 +505,6 @@ let ring_equality env evd (r,add,mul,opp,req) = | None -> plapp evd coq_eq_smorph [|r;add;mul|] in let setoid = Typing.e_solve_evars env evd setoid in let op_morph = Typing.e_solve_evars env evd op_morph in - let setoid = EConstr.of_constr setoid in - let op_morph = EConstr.of_constr op_morph in (setoid,op_morph) | _ -> let setoid = setoid_of_relation (Global.env ()) evd r req in @@ -594,6 +592,7 @@ let make_hyp_list env evd lH = (plapp evd coq_nil [|carrier|]) in let l' = Typing.e_solve_evars env evd l in + let l' = EConstr.Unsafe.to_constr l' in Evarutil.nf_evars_universes !evd l' let interp_power env evd pow = @@ -756,7 +755,7 @@ let ring_lookup (f : Value.t) lH rl t = let rl = make_args_list sigma rl t in let evdref = ref sigma in let e = find_ring_structure env sigma rl in - let rl = carg (make_term_list env evdref (EConstr.of_constr e.ring_carrier) rl) in + let rl = Value.of_constr (make_term_list env evdref (EConstr.of_constr e.ring_carrier) rl) in let lH = carg (make_hyp_list env evdref lH) in let ring = ltac_ring_structure e in Proofview.tclTHEN (Proofview.Unsafe.tclEVARS !evdref) (ltac_apply f (ring@[lH;rl])) @@ -1039,7 +1038,7 @@ let field_lookup (f : Value.t) lH rl t = let rl = make_args_list sigma rl t in let evdref = ref sigma in let e = find_field_structure env sigma rl in - let rl = carg (make_term_list env evdref (EConstr.of_constr e.field_carrier) rl) in + let rl = Value.of_constr (make_term_list env evdref (EConstr.of_constr e.field_carrier) rl) in let lH = carg (make_hyp_list env evdref lH) in let field = ltac_field_structure e in Proofview.tclTHEN (Proofview.Unsafe.tclEVARS !evdref) (ltac_apply f (field@[lH;rl])) diff --git a/plugins/ssrmatching/ssrmatching.ml4 b/plugins/ssrmatching/ssrmatching.ml4 index 9dcc6c4cc..eb5f401e3 100644 --- a/plugins/ssrmatching/ssrmatching.ml4 +++ b/plugins/ssrmatching/ssrmatching.ml4 @@ -1353,6 +1353,7 @@ let fill_occ_term env cl occ sigma0 (sigma, t) = let pf_fill_occ_term gl occ t = let sigma0 = project gl and env = pf_env gl and concl = pf_concl gl in + let concl = EConstr.Unsafe.to_constr concl in let cl,(_,t) = fill_occ_term env concl occ sigma0 t in cl, t @@ -1388,6 +1389,7 @@ let ssrpatterntac _ist (arg_ist,arg) gl = let pat = interp_rpattern arg_ist gl arg in let sigma0 = project gl in let concl0 = pf_concl gl in + let concl0 = EConstr.Unsafe.to_constr concl0 in let (t, uc), concl_x = fill_occ_pattern (Global.env()) sigma0 concl0 pat noindex 1 in let t = EConstr.of_constr t in @@ -1416,6 +1418,7 @@ let ssrinstancesof ist arg gl = let ok rhs lhs ise = true in (* not (Term.eq_constr lhs (Evarutil.nf_evar ise rhs)) in *) let env, sigma, concl = pf_env gl, project gl, pf_concl gl in + let concl = EConstr.Unsafe.to_constr concl in let sigma0, cpat = interp_cpattern ist gl arg None in let pat = match cpat with T x -> x | _ -> errorstrm (str"Not supported") in let etpat, tpat = mk_tpattern env sigma (sigma0,pat) (ok pat) L2R pat in diff --git a/pretyping/cases.ml b/pretyping/cases.ml index 565a9725c..eea94f021 100644 --- a/pretyping/cases.ml +++ b/pretyping/cases.ml @@ -2076,7 +2076,7 @@ let constr_of_pat env evdref arsign pat avoid = let IndType (indf, _) = try find_rectype env ( !evdref) (lift (-(List.length realargs)) ty) with Not_found -> error_case_not_inductive env !evdref - {uj_val = ty; uj_type = EConstr.of_constr (Typing.unsafe_type_of env !evdref ty)} + {uj_val = ty; uj_type = Typing.unsafe_type_of env !evdref ty} in let (ind,u), params = dest_ind_family indf in let params = List.map EConstr.of_constr params in diff --git a/pretyping/cbv.ml b/pretyping/cbv.ml index a42061f28..e18625c42 100644 --- a/pretyping/cbv.ml +++ b/pretyping/cbv.ml @@ -377,7 +377,7 @@ and cbv_norm_value info = function (* reduction under binders *) (* with profiling *) let cbv_norm infos constr = let constr = EConstr.Unsafe.to_constr constr in - with_stats (lazy (cbv_norm_term infos (subs_id 0) constr)) + EConstr.of_constr (with_stats (lazy (cbv_norm_term infos (subs_id 0) constr))) type cbv_infos = cbv_value infos diff --git a/pretyping/cbv.mli b/pretyping/cbv.mli index 3d1745767..b014af2c7 100644 --- a/pretyping/cbv.mli +++ b/pretyping/cbv.mli @@ -8,6 +8,7 @@ open Names open Term +open EConstr open Environ open CClosure open Esubst @@ -19,10 +20,13 @@ open Esubst type cbv_infos val create_cbv_infos : RedFlags.reds -> env -> Evd.evar_map -> cbv_infos -val cbv_norm : cbv_infos -> EConstr.constr -> constr +val cbv_norm : cbv_infos -> constr -> constr (*********************************************************************** i This is for cbv debug *) + +open Term + type cbv_value = | VAL of int * constr | STACK of int * cbv_value * cbv_stack diff --git a/pretyping/coercion.ml b/pretyping/coercion.ml index f569d9fc4..ead44cee2 100644 --- a/pretyping/coercion.ml +++ b/pretyping/coercion.ml @@ -194,7 +194,7 @@ and coerce loc env evdref (x : EConstr.constr) (y : EConstr.constr) (subst1 hdy restT') (succ i) (fun x -> eq_app (co x)) else Some (fun x -> let term = co x in - EConstr.of_constr (Typing.e_solve_evars env evdref term)) + Typing.e_solve_evars env evdref term) in if isEvar !evdref c || isEvar !evdref c' || not (Program.is_program_generalized_coercion ()) then (* Second-order unification needed. *) @@ -302,7 +302,7 @@ and coerce loc env evdref (x : EConstr.constr) (y : EConstr.constr) with NoSubtacCoercion -> let typ = Typing.unsafe_type_of env evm c in let typ' = Typing.unsafe_type_of env evm c' in - coerce_application (EConstr.of_constr typ) (EConstr.of_constr typ') c c' l l') + coerce_application typ typ' c c' l l') else subco () | x, y when EConstr.eq_constr !evdref c c' -> @@ -310,7 +310,7 @@ and coerce loc env evdref (x : EConstr.constr) (y : EConstr.constr) let evm = !evdref in let lam_type = Typing.unsafe_type_of env evm c in let lam_type' = Typing.unsafe_type_of env evm c' in - coerce_application (EConstr.of_constr lam_type) (EConstr.of_constr lam_type') c c' l l' + coerce_application lam_type lam_type' c c' l l' else subco () | _ -> subco ()) | _, _ -> subco () @@ -341,7 +341,7 @@ let app_coercion env evdref coercion v = | None -> v | Some f -> let v' = Typing.e_solve_evars env evdref (f v) in - whd_betaiota !evdref (EConstr.of_constr v') + whd_betaiota !evdref v' let coerce_itf loc env evd v t c1 = let evdref = ref evd in diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml index 09b99983c..f76f608d0 100644 --- a/pretyping/pretyping.ml +++ b/pretyping/pretyping.ml @@ -524,7 +524,6 @@ let pretype_ref loc evdref env ref us = let evd, c = pretype_global loc univ_flexible env !evdref ref us in let () = evdref := evd in let ty = unsafe_type_of env.ExtraEnv.env evd c in - let ty = EConstr.of_constr ty in make_judge c ty let judge_of_Type loc evd s = @@ -1194,16 +1193,16 @@ let understand let understand_tcc ?(flags=all_no_fail_flags) env sigma ?(expected_type=WithoutTypeConstraint) c = let (sigma, c) = ise_pretype_gen flags env sigma empty_lvar expected_type c in - (sigma, EConstr.Unsafe.to_constr c) + (sigma, c) let understand_tcc_evars ?(flags=all_no_fail_flags) env evdref ?(expected_type=WithoutTypeConstraint) c = let sigma, c = ise_pretype_gen flags env !evdref empty_lvar expected_type c in evdref := sigma; - EConstr.Unsafe.to_constr c + c let understand_ltac flags env sigma lvar kind c = let (sigma, c) = ise_pretype_gen flags env sigma lvar kind c in - (sigma, EConstr.Unsafe.to_constr c) + (sigma, c) let constr_flags = { use_typeclasses = true; @@ -1225,7 +1224,6 @@ let type_uconstr ?(flags = constr_flags) } in let sigma = Sigma.to_evar_map sigma in let (sigma, c) = understand_ltac flags env sigma vars expected_type term in - let c = EConstr.of_constr c in Sigma.Unsafe.of_pair (c, sigma) end } diff --git a/pretyping/pretyping.mli b/pretyping/pretyping.mli index a1602088a..825d73f1f 100644 --- a/pretyping/pretyping.mli +++ b/pretyping/pretyping.mli @@ -76,10 +76,10 @@ val allow_anonymous_refs : bool ref evar_map is modified explicitly or by side-effect. *) val understand_tcc : ?flags:inference_flags -> env -> evar_map -> - ?expected_type:typing_constraint -> glob_constr -> open_constr + ?expected_type:typing_constraint -> glob_constr -> evar_map * EConstr.constr val understand_tcc_evars : ?flags:inference_flags -> env -> evar_map ref -> - ?expected_type:typing_constraint -> glob_constr -> constr + ?expected_type:typing_constraint -> glob_constr -> EConstr.constr (** More general entry point with evars from ltac *) @@ -95,7 +95,7 @@ val understand_tcc_evars : ?flags:inference_flags -> env -> evar_map ref -> val understand_ltac : inference_flags -> env -> evar_map -> ltac_var_map -> - typing_constraint -> glob_constr -> pure_open_constr + typing_constraint -> glob_constr -> evar_map * EConstr.constr (** Standard call to get a constr from a glob_constr, resolving implicit arguments and coercions, and compiling pattern-matching; diff --git a/pretyping/tacred.ml b/pretyping/tacred.ml index 3fc01c86c..2b496f926 100644 --- a/pretyping/tacred.ml +++ b/pretyping/tacred.ml @@ -1133,7 +1133,7 @@ let fold_commands cl env sigma c = (* call by value reduction functions *) let cbv_norm_flags flags env sigma t = - EConstr.of_constr (cbv_norm (create_cbv_infos flags env sigma) t) + cbv_norm (create_cbv_infos flags env sigma) t let cbv_beta = cbv_norm_flags beta empty_env let cbv_betaiota = cbv_norm_flags betaiota empty_env diff --git a/pretyping/typing.ml b/pretyping/typing.ml index d24160ea5..7baff421f 100644 --- a/pretyping/typing.ml +++ b/pretyping/typing.ml @@ -376,7 +376,7 @@ let unsafe_type_of env evd c = let evdref = ref evd in let env = enrich_env env evdref in let j = execute env evdref c in - EConstr.Unsafe.to_constr j.uj_type + j.uj_type (* Sort of a type *) @@ -411,6 +411,6 @@ let e_solve_evars env evdref c = let env = enrich_env env evdref in let c = (execute env evdref c).uj_val in (* side-effect on evdref *) - nf_evar !evdref (EConstr.Unsafe.to_constr c) + EConstr.of_constr (nf_evar !evdref (EConstr.Unsafe.to_constr c)) -let _ = Evarconv.set_solve_evars (fun env evdref c -> EConstr.of_constr (e_solve_evars env evdref c)) +let _ = Evarconv.set_solve_evars (fun env evdref c -> e_solve_evars env evdref c) diff --git a/pretyping/typing.mli b/pretyping/typing.mli index bf26358a2..91134b499 100644 --- a/pretyping/typing.mli +++ b/pretyping/typing.mli @@ -9,46 +9,47 @@ open Names open Term open Environ +open EConstr open Evd (** This module provides the typing machine with existential variables and universes. *) (** Typecheck a term and return its type. May trigger an evarmap leak. *) -val unsafe_type_of : env -> evar_map -> EConstr.constr -> types +val unsafe_type_of : env -> evar_map -> constr -> types (** Typecheck a term and return its type + updated evars, optionally refreshing universes *) -val type_of : ?refresh:bool -> env -> evar_map -> EConstr.constr -> evar_map * EConstr.types +val type_of : ?refresh:bool -> env -> evar_map -> constr -> evar_map * types (** Variant of [type_of] using references instead of state-passing. *) -val e_type_of : ?refresh:bool -> env -> evar_map ref -> EConstr.constr -> EConstr.types +val e_type_of : ?refresh:bool -> env -> evar_map ref -> constr -> types (** Typecheck a type and return its sort *) -val e_sort_of : env -> evar_map ref -> EConstr.types -> sorts +val e_sort_of : env -> evar_map ref -> types -> sorts (** Typecheck a term has a given type (assuming the type is OK) *) -val e_check : env -> evar_map ref -> EConstr.constr -> EConstr.types -> unit +val e_check : env -> evar_map ref -> constr -> types -> unit (** Returns the instantiated type of a metavariable *) -val meta_type : evar_map -> metavariable -> EConstr.types +val meta_type : evar_map -> metavariable -> types (** Solve existential variables using typing *) -val e_solve_evars : env -> evar_map ref -> EConstr.constr -> constr +val e_solve_evars : env -> evar_map ref -> constr -> constr (** Raise an error message if incorrect elimination for this inductive *) (** (first constr is term to match, second is return predicate) *) -val check_allowed_sort : env -> evar_map -> pinductive -> EConstr.constr -> EConstr.constr -> +val check_allowed_sort : env -> evar_map -> pinductive -> constr -> constr -> unit (** Raise an error message if bodies have types not unifiable with the expected ones *) val check_type_fixpoint : Loc.t -> env -> evar_map ref -> - Names.Name.t array -> EConstr.types array -> EConstr.unsafe_judgment array -> unit + Names.Name.t array -> types array -> unsafe_judgment array -> unit -val judge_of_prop : EConstr.unsafe_judgment -val judge_of_set : EConstr.unsafe_judgment +val judge_of_prop : unsafe_judgment +val judge_of_set : unsafe_judgment val judge_of_abstraction : Environ.env -> Name.t -> - EConstr.unsafe_type_judgment -> EConstr.unsafe_judgment -> EConstr.unsafe_judgment + unsafe_type_judgment -> unsafe_judgment -> unsafe_judgment val judge_of_product : Environ.env -> Name.t -> - EConstr.unsafe_type_judgment -> EConstr.unsafe_type_judgment -> EConstr.unsafe_judgment + unsafe_type_judgment -> unsafe_type_judgment -> unsafe_judgment diff --git a/pretyping/unification.ml b/pretyping/unification.ml index c6fad1a34..5bb865310 100644 --- a/pretyping/unification.ml +++ b/pretyping/unification.ml @@ -1228,7 +1228,7 @@ let applyHead env (type r) (evd : r Sigma.t) n c = apprec (n-1) (mkApp(c,[|evar|])) (subst1 evar c2) (p +> q) evd' | _ -> error "Apply_Head_Then" in - apprec n c (EConstr.of_constr (Typing.unsafe_type_of env (Sigma.to_evar_map evd) c)) Sigma.refl evd + apprec n c (Typing.unsafe_type_of env (Sigma.to_evar_map evd) c) Sigma.refl evd let is_mimick_head sigma ts f = match EConstr.kind sigma f with diff --git a/proofs/clenv.ml b/proofs/clenv.ml index 7269c61e3..e580bccc3 100644 --- a/proofs/clenv.ml +++ b/proofs/clenv.ml @@ -133,7 +133,7 @@ let mk_clenv_from_n gls n (c,cty) = let mk_clenv_from gls = mk_clenv_from_n gls None -let mk_clenv_type_of gls t = mk_clenv_from gls (t,EConstr.of_constr (pf_type_of gls t)) +let mk_clenv_type_of gls t = mk_clenv_from gls (t,pf_type_of gls t) (******************************************************************) diff --git a/proofs/evar_refiner.ml b/proofs/evar_refiner.ml index 0d65faf12..1ddb67edd 100644 --- a/proofs/evar_refiner.ml +++ b/proofs/evar_refiner.ml @@ -27,9 +27,9 @@ let depends_on_evar sigma evk _ (pbty,_,t1,t2) = with NoHeadEvar -> false let define_and_solve_constraints evk c env evd = - if Termops.occur_evar evd evk (EConstr.of_constr c) then - Pretype_errors.error_occur_check env evd evk (EConstr.of_constr c); - let evd = define evk c evd in + if Termops.occur_evar evd evk c then + Pretype_errors.error_occur_check env evd evk c; + let evd = define evk (EConstr.Unsafe.to_constr c) evd in let (evd,pbs) = extract_changed_conv_pbs evd (depends_on_evar evd evk) in match List.fold_left diff --git a/proofs/logic.ml b/proofs/logic.ml index 8b890f6f8..59ce8ffa6 100644 --- a/proofs/logic.ml +++ b/proofs/logic.ml @@ -330,7 +330,7 @@ let meta_free_prefix sigma a = with Stop acc -> Array.rev_of_list acc let goal_type_of env sigma c = - if !check then unsafe_type_of env sigma (EConstr.of_constr c) + if !check then EConstr.Unsafe.to_constr (unsafe_type_of env sigma (EConstr.of_constr c)) else EConstr.Unsafe.to_constr (Retyping.get_type_of env sigma (EConstr.of_constr c)) let rec mk_refgoals sigma goal goalacc conclty trm = diff --git a/proofs/tacmach.ml b/proofs/tacmach.ml index 3641ad74d..2fab026ea 100644 --- a/proofs/tacmach.ml +++ b/proofs/tacmach.ml @@ -46,11 +46,11 @@ let test_conversion cb env sigma c1 c2 = let c2 = EConstr.Unsafe.to_constr c2 in test_conversion cb env sigma c1 c2 -let pf_concl gls = EConstr.Unsafe.to_constr (Goal.V82.concl (project gls) (sig_it gls)) +let pf_concl gls = Goal.V82.concl (project gls) (sig_it gls) let pf_hyps_types gls = let sign = Environ.named_context (pf_env gls) in List.map (function LocalAssum (id,x) - | LocalDef (id,_,x) -> id, x) + | LocalDef (id,_,x) -> id, EConstr.of_constr x) sign let pf_nth_hyp_id gls n = List.nth (pf_hyps gls) (n-1) |> NamedDecl.get_id @@ -64,7 +64,7 @@ let pf_get_hyp gls id = raise (RefinerError (NoSuchHyp id)) let pf_get_hyp_typ gls id = - id |> pf_get_hyp gls |> NamedDecl.get_type + id |> pf_get_hyp gls |> NamedDecl.get_type |> EConstr.of_constr let pf_ids_of_hyps gls = ids_of_named_context (pf_hyps gls) @@ -77,7 +77,7 @@ let pf_get_new_ids ids gls = (fun id acc -> (next_ident_away id (acc@avoid))::acc) ids [] -let pf_global gls id = Constrintern.construct_reference (pf_hyps gls) id +let pf_global gls id = EConstr.of_constr (Constrintern.construct_reference (pf_hyps gls) id) let pf_reduction_of_red_expr gls re c = let (redfun, _) = reduction_of_red_expr (pf_env gls) re in @@ -103,7 +103,7 @@ let pf_get_type_of = pf_reduce Retyping.get_type_of let pf_conv_x gl = pf_reduce test_conversion gl Reduction.CONV let pf_conv_x_leq gl = pf_reduce test_conversion gl Reduction.CUMUL -let pf_const_value = pf_reduce (fun env _ -> constant_value_in env) +let pf_const_value = pf_reduce (fun env _ c -> EConstr.of_constr (constant_value_in env c)) let pf_reduce_to_quantified_ind = pf_reduce reduce_to_quantified_ind let pf_reduce_to_atomic_ind = pf_reduce reduce_to_atomic_ind @@ -173,7 +173,7 @@ module New = struct (** We only check for the existence of an [id] in [hyps] *) let gl = Proofview.Goal.assume gl in let hyps = Proofview.Goal.hyps gl in - Constrintern.construct_reference hyps id + EConstr.of_constr (Constrintern.construct_reference hyps id) let pf_env = Proofview.Goal.env let pf_concl = Proofview.Goal.concl @@ -205,13 +205,13 @@ module New = struct sign let pf_get_hyp_typ id gl = - pf_get_hyp id gl |> NamedDecl.get_type + pf_get_hyp id gl |> NamedDecl.get_type |> EConstr.of_constr let pf_hyps_types gl = let env = Proofview.Goal.env gl in let sign = Environ.named_context env in List.map (function LocalAssum (id,x) - | LocalDef (id,_,x) -> id, x) + | LocalDef (id,_,x) -> id, EConstr.of_constr x) sign let pf_last_hyp gl = @@ -241,6 +241,6 @@ module New = struct let pf_whd_all gl t = pf_apply whd_all gl t let pf_compute gl t = pf_apply compute gl t - let pf_nf_evar gl t = nf_evar (project gl) t + let pf_nf_evar gl t = EConstr.of_constr (nf_evar (project gl) (EConstr.Unsafe.to_constr t)) end diff --git a/proofs/tacmach.mli b/proofs/tacmach.mli index f0471bf66..efc3dbf55 100644 --- a/proofs/tacmach.mli +++ b/proofs/tacmach.mli @@ -9,6 +9,7 @@ open Names open Term open Environ +open EConstr open Evd open Proof_type open Redexpr @@ -40,9 +41,9 @@ val pf_nth_hyp_id : goal sigma -> int -> Id.t val pf_last_hyp : goal sigma -> Context.Named.Declaration.t val pf_ids_of_hyps : goal sigma -> Id.t list val pf_global : goal sigma -> Id.t -> constr -val pf_unsafe_type_of : goal sigma -> EConstr.constr -> types -val pf_type_of : goal sigma -> EConstr.constr -> evar_map * EConstr.types -val pf_hnf_type_of : goal sigma -> EConstr.constr -> EConstr.types +val pf_unsafe_type_of : goal sigma -> constr -> types +val pf_type_of : goal sigma -> constr -> evar_map * types +val pf_hnf_type_of : goal sigma -> constr -> types val pf_get_hyp : goal sigma -> Id.t -> Context.Named.Declaration.t val pf_get_hyp_typ : goal sigma -> Id.t -> types @@ -50,7 +51,7 @@ val pf_get_hyp_typ : goal sigma -> Id.t -> types val pf_get_new_id : Id.t -> goal sigma -> Id.t val pf_get_new_ids : Id.t list -> goal sigma -> Id.t list -val pf_reduction_of_red_expr : goal sigma -> red_expr -> EConstr.constr -> evar_map * EConstr.constr +val pf_reduction_of_red_expr : goal sigma -> red_expr -> constr -> evar_map * constr val pf_apply : (env -> evar_map -> 'a) -> goal sigma -> 'a @@ -63,35 +64,35 @@ val pf_e_reduce : (env -> evar_map -> constr -> evar_map * constr) -> goal sigma -> constr -> evar_map * constr -val pf_whd_all : goal sigma -> EConstr.constr -> EConstr.constr -val pf_hnf_constr : goal sigma -> EConstr.constr -> EConstr.constr -val pf_nf : goal sigma -> EConstr.constr -> EConstr.constr -val pf_nf_betaiota : goal sigma -> EConstr.constr -> EConstr.constr -val pf_reduce_to_quantified_ind : goal sigma -> EConstr.types -> pinductive * EConstr.types -val pf_reduce_to_atomic_ind : goal sigma -> EConstr.types -> pinductive * EConstr.types -val pf_compute : goal sigma -> EConstr.constr -> EConstr.constr +val pf_whd_all : goal sigma -> constr -> constr +val pf_hnf_constr : goal sigma -> constr -> constr +val pf_nf : goal sigma -> constr -> constr +val pf_nf_betaiota : goal sigma -> constr -> constr +val pf_reduce_to_quantified_ind : goal sigma -> types -> pinductive * types +val pf_reduce_to_atomic_ind : goal sigma -> types -> pinductive * types +val pf_compute : goal sigma -> constr -> constr val pf_unfoldn : (occurrences * evaluable_global_reference) list - -> goal sigma -> EConstr.constr -> EConstr.constr + -> goal sigma -> constr -> constr val pf_const_value : goal sigma -> pconstant -> constr -val pf_conv_x : goal sigma -> EConstr.constr -> EConstr.constr -> bool -val pf_conv_x_leq : goal sigma -> EConstr.constr -> EConstr.constr -> bool +val pf_conv_x : goal sigma -> constr -> constr -> bool +val pf_conv_x_leq : goal sigma -> constr -> constr -> bool -val pf_matches : goal sigma -> constr_pattern -> EConstr.constr -> patvar_map -val pf_is_matching : goal sigma -> constr_pattern -> EConstr.constr -> bool +val pf_matches : goal sigma -> constr_pattern -> constr -> patvar_map +val pf_is_matching : goal sigma -> constr_pattern -> constr -> bool (** {6 The most primitive tactics. } *) val refiner : rule -> tactic -val internal_cut_no_check : bool -> Id.t -> EConstr.types -> tactic -val refine_no_check : EConstr.constr -> tactic +val internal_cut_no_check : bool -> Id.t -> types -> tactic +val refine_no_check : constr -> tactic (** {6 The most primitive tactics with consistency and type checking } *) -val internal_cut : bool -> Id.t -> EConstr.types -> tactic -val internal_cut_rev : bool -> Id.t -> EConstr.types -> tactic -val refine : EConstr.constr -> tactic +val internal_cut : bool -> Id.t -> types -> tactic +val internal_cut_rev : bool -> Id.t -> types -> tactic +val refine : constr -> tactic (** {6 Pretty-printing functions (debug only). } *) val pr_gls : goal sigma -> Pp.std_ppcmds @@ -106,11 +107,11 @@ module New : sig val project : ('a, 'r) Proofview.Goal.t -> Evd.evar_map val pf_env : ('a, 'r) Proofview.Goal.t -> Environ.env - val pf_concl : ([ `NF ], 'r) Proofview.Goal.t -> EConstr.types + val pf_concl : ([ `NF ], 'r) Proofview.Goal.t -> types - val pf_unsafe_type_of : ('a, 'r) Proofview.Goal.t -> EConstr.constr -> Term.types - val pf_type_of : ('a, 'r) Proofview.Goal.t -> EConstr.constr -> evar_map * EConstr.types - val pf_conv_x : ('a, 'r) Proofview.Goal.t -> EConstr.t -> EConstr.t -> bool + val pf_unsafe_type_of : ('a, 'r) Proofview.Goal.t -> constr -> types + val pf_type_of : ('a, 'r) Proofview.Goal.t -> constr -> evar_map * types + val pf_conv_x : ('a, 'r) Proofview.Goal.t -> t -> t -> bool val pf_get_new_id : identifier -> ([ `NF ], 'r) Proofview.Goal.t -> identifier val pf_ids_of_hyps : ('a, 'r) Proofview.Goal.t -> identifier list @@ -120,16 +121,16 @@ module New : sig val pf_get_hyp_typ : identifier -> ([ `NF ], 'r) Proofview.Goal.t -> types val pf_last_hyp : ([ `NF ], 'r) Proofview.Goal.t -> Context.Named.Declaration.t - val pf_nf_concl : ([ `LZ ], 'r) Proofview.Goal.t -> EConstr.types - val pf_reduce_to_quantified_ind : ('a, 'r) Proofview.Goal.t -> EConstr.types -> pinductive * EConstr.types + val pf_nf_concl : ([ `LZ ], 'r) Proofview.Goal.t -> types + val pf_reduce_to_quantified_ind : ('a, 'r) Proofview.Goal.t -> types -> pinductive * types - val pf_hnf_constr : ('a, 'r) Proofview.Goal.t -> EConstr.constr -> EConstr.types - val pf_hnf_type_of : ('a, 'r) Proofview.Goal.t -> EConstr.constr -> EConstr.types + val pf_hnf_constr : ('a, 'r) Proofview.Goal.t -> constr -> types + val pf_hnf_type_of : ('a, 'r) Proofview.Goal.t -> constr -> types - val pf_whd_all : ('a, 'r) Proofview.Goal.t -> EConstr.t -> EConstr.constr - val pf_compute : ('a, 'r) Proofview.Goal.t -> EConstr.t -> EConstr.constr + val pf_whd_all : ('a, 'r) Proofview.Goal.t -> constr -> constr + val pf_compute : ('a, 'r) Proofview.Goal.t -> constr -> constr - val pf_matches : ('a, 'r) Proofview.Goal.t -> constr_pattern -> EConstr.constr -> patvar_map + val pf_matches : ('a, 'r) Proofview.Goal.t -> constr_pattern -> constr -> patvar_map val pf_nf_evar : ('a, 'r) Proofview.Goal.t -> constr -> constr diff --git a/stm/lemmas.ml b/stm/lemmas.ml index 9942e911a..e3f1c1ac2 100644 --- a/stm/lemmas.ml +++ b/stm/lemmas.ml @@ -461,6 +461,7 @@ let start_proof_com ?inference_hook kind thms hook = let thms = List.map (fun (sopt,(bl,t,guard)) -> let impls, ((env, ctx), imps) = interp_context_evars env0 evdref bl in let t', imps' = interp_type_evars_impls ~impls env evdref t in + let t' = EConstr.Unsafe.to_constr t' in let flags = all_and_fail_flags in let flags = { flags with use_hook = inference_hook } in evdref := solve_remaining_evars flags env !evdref (Evd.empty,!evdref); diff --git a/tactics/autorewrite.ml b/tactics/autorewrite.ml index f2e98ee01..f43f4b250 100644 --- a/tactics/autorewrite.ml +++ b/tactics/autorewrite.ml @@ -257,7 +257,7 @@ type hypinfo = { let decompose_applied_relation metas env sigma c ctype left2right = let find_rel ty = - let eqclause = Clenv.mk_clenv_from_env env sigma None (EConstr.of_constr c,EConstr.of_constr ty) in + let eqclause = Clenv.mk_clenv_from_env env sigma None (EConstr.of_constr c,ty) in let eqclause = if metas then eqclause else clenv_pose_metas_as_evars eqclause (Evd.undefined_metas eqclause.evd) @@ -274,6 +274,8 @@ let decompose_applied_relation metas env sigma c ctype left2right = let ty1, ty2 = Typing.unsafe_type_of env eqclause.evd (EConstr.of_constr c1), Typing.unsafe_type_of env eqclause.evd (EConstr.of_constr c2) in + let ty = EConstr.Unsafe.to_constr ty in + let ty1 = EConstr.Unsafe.to_constr ty1 in (* if not (evd_convertible env eqclause.evd ty1 ty2) then None *) (* else *) Some { hyp_cl=eqclause; hyp_prf=EConstr.Unsafe.to_constr (Clenv.clenv_value eqclause); hyp_ty = ty; @@ -284,9 +286,8 @@ let decompose_applied_relation metas env sigma c ctype left2right = match find_rel ctype with | Some c -> Some c | None -> - let ctx,t' = Reductionops.splay_prod_assum env sigma (EConstr.of_constr ctype) in (* Search for underlying eq *) - let t' = EConstr.Unsafe.to_constr t' in - match find_rel (Term.it_mkProd_or_LetIn t' ctx) with + let ctx,t' = Reductionops.splay_prod_assum env sigma ctype in (* Search for underlying eq *) + match find_rel (it_mkProd_or_LetIn t' ctx) with | Some c -> Some c | None -> None @@ -296,7 +297,7 @@ let find_applied_relation metas loc env sigma c left2right = | Some c -> c | None -> user_err ~loc ~hdr:"decompose_applied_relation" - (str"The type" ++ spc () ++ Printer.pr_constr_env env sigma ctype ++ + (str"The type" ++ spc () ++ Printer.pr_econstr_env env sigma ctype ++ spc () ++ str"of this term does not end with an applied relation.") (* To add rewriting rules to a base *) diff --git a/tactics/class_tactics.ml b/tactics/class_tactics.ml index 84ca0aa8f..fa2c21ac3 100644 --- a/tactics/class_tactics.ml +++ b/tactics/class_tactics.ml @@ -227,7 +227,6 @@ let e_give_exact flags poly (c,clenv) gl = else c, gl in let t1 = pf_unsafe_type_of gl c in - let t1 = EConstr.of_constr t1 in Proofview.V82.of_tactic (Clenvtac.unify ~flags t1 <*> exact_no_check c) gl let unify_e_resolve poly flags = { enter = begin fun gls (c,_,clenv) -> @@ -1515,7 +1514,6 @@ let autoapply c i gl = let flags = auto_unif_flags Evar.Set.empty (Hints.Hint_db.transparent_state (Hints.searchtable_map i)) in let cty = pf_unsafe_type_of gl c in - let cty = EConstr.of_constr cty in let ce = mk_clenv_from gl (c,cty) in let tac = { enter = fun gl -> (unify_e_resolve false flags).enter gl ((c,cty,Univ.ContextSet.empty),0,ce) } in diff --git a/tactics/contradiction.ml b/tactics/contradiction.ml index afc7e1547..a3a448aad 100644 --- a/tactics/contradiction.ml +++ b/tactics/contradiction.ml @@ -116,7 +116,6 @@ let contradiction_term (c,lbind as cl) = let env = Proofview.Goal.env gl in let type_of = Tacmach.New.pf_unsafe_type_of gl in let typ = type_of c in - let typ = EConstr.of_constr typ in let _, ccl = splay_prod env sigma typ in if is_empty_type sigma ccl then Tacticals.New.tclTHEN diff --git a/tactics/eauto.ml b/tactics/eauto.ml index 92e59c5ce..01f21910c 100644 --- a/tactics/eauto.ml +++ b/tactics/eauto.ml @@ -32,7 +32,6 @@ let eauto_unif_flags = auto_flags_of_state full_transparent_state let e_give_exact ?(flags=eauto_unif_flags) c = Proofview.Goal.enter { enter = begin fun gl -> let t1 = Tacmach.New.pf_unsafe_type_of gl c in - let t1 = EConstr.of_constr t1 in let t2 = Tacmach.New.pf_concl (Proofview.Goal.assume gl) in let sigma = Tacmach.New.project gl in if occur_existential sigma t1 || occur_existential sigma t2 then @@ -290,7 +289,7 @@ module SearchProblem = struct in let rec_tacs = let l = - let concl = Reductionops.nf_evar (project g)(pf_concl g) in + let concl = Reductionops.nf_evar (project g) (EConstr.Unsafe.to_constr (pf_concl g)) in let concl = EConstr.of_constr concl in filter_tactics s.tacres (e_possible_resolve (project g) s.dblist (List.hd s.localdb) secvars concl) @@ -516,7 +515,7 @@ let autounfold_one db cl = (Id.Set.union ids i, Cset.union csts c)) (Id.Set.empty, Cset.empty) db in let did, c' = unfold_head env sigma st - (match cl with Some (id, _) -> EConstr.of_constr (Tacmach.New.pf_get_hyp_typ id gl) | None -> concl) + (match cl with Some (id, _) -> Tacmach.New.pf_get_hyp_typ id gl | None -> concl) in if did then match cl with diff --git a/tactics/elim.ml b/tactics/elim.ml index ef848c2e1..a4158f821 100644 --- a/tactics/elim.ml +++ b/tactics/elim.ml @@ -82,7 +82,6 @@ let general_decompose recognizer c = let type_of = pf_unsafe_type_of gl in let sigma = project gl in let typc = type_of c in - let typc = EConstr.of_constr typc in tclTHENS (cut typc) [ tclTHEN (intro_using tmphyp_name) (onLastHypId @@ -136,7 +135,6 @@ let induction_trailer abs_i abs_j bargs = (fun id -> Proofview.Goal.nf_enter { enter = begin fun gl -> let idty = pf_unsafe_type_of gl (mkVar id) in - let idty = EConstr.of_constr idty in let fvty = global_vars (pf_env gl) (project gl) idty in let possible_bring_hyps = (List.tl (nLastDecls gl (abs_j - abs_i))) @ bargs.Tacticals.assums diff --git a/tactics/eqdecide.ml b/tactics/eqdecide.ml index 16e0d9684..df60f2c66 100644 --- a/tactics/eqdecide.ml +++ b/tactics/eqdecide.ml @@ -163,7 +163,6 @@ let rec solveArg hyps eqonleft op largs rargs = match largs, rargs with | a1 :: largs, a2 :: rargs -> Proofview.Goal.enter { enter = begin fun gl -> let rectype = pf_unsafe_type_of gl a1 in - let rectype = EConstr.of_constr rectype in let decide = mkDecideEqGoal eqonleft op rectype a1 a2 in let tac hyp = solveArg (hyp :: hyps) eqonleft op largs rargs in let subtacs = @@ -236,7 +235,6 @@ let decideEquality rectype = let compare c1 c2 = Proofview.Goal.enter { enter = begin fun gl -> let rectype = pf_unsafe_type_of gl c1 in - let rectype = EConstr.of_constr rectype in let decide = mkDecideEqGoal true (build_coq_sumbool ()) rectype c1 c2 in (tclTHENS (cut decide) [(tclTHEN intro diff --git a/tactics/equality.ml b/tactics/equality.ml index e1c39bb34..7dcfd419e 100644 --- a/tactics/equality.ml +++ b/tactics/equality.ml @@ -311,7 +311,7 @@ let general_elim_clause with_evars frzevars tac cls c t l l2r elim = in let typ = match cls with | None -> pf_nf_concl gl - | Some id -> EConstr.of_constr (pf_get_hyp_typ id (Proofview.Goal.assume gl)) + | Some id -> pf_get_hyp_typ id (Proofview.Goal.assume gl) in let cs = instantiate_lemma typ in if firstonly then tclFIRST (List.map try_clause cs) @@ -407,7 +407,7 @@ let find_elim hdcncl lft2rgt dep cls ot gl = let type_of_clause cls gl = match cls with | None -> Proofview.Goal.concl gl - | Some id -> EConstr.of_constr (pf_get_hyp_typ id gl) + | Some id -> pf_get_hyp_typ id gl let leibniz_rewrite_ebindings_clause cls lft2rgt tac c t l with_evars frzevars dep_proof_ok hdcncl = Proofview.Goal.nf_s_enter { s_enter = begin fun gl -> @@ -950,7 +950,6 @@ let gen_absurdity id = Proofview.Goal.enter { enter = begin fun gl -> let sigma = project gl in let hyp_typ = pf_get_hyp_typ id (Proofview.Goal.assume gl) in - let hyp_typ = EConstr.of_constr hyp_typ in if is_empty_type sigma hyp_typ then simplest_elim (mkVar id) @@ -1027,7 +1026,6 @@ let onEquality with_evars tac (c,lbindc) = let type_of = pf_unsafe_type_of gl in let reduce_to_quantified_ind = pf_apply Tacred.reduce_to_quantified_ind gl in let t = type_of c in - let t = EConstr.of_constr t in let t' = try snd (reduce_to_quantified_ind t) with UserError _ -> t in let eq_clause = pf_apply make_clenv_binding gl (c,t') lbindc in let eq_clause' = Clenvtac.clenv_pose_dependent_evars with_evars eq_clause in @@ -1136,7 +1134,7 @@ let minimal_free_rels_rec env sigma = let rec minimalrec_free_rels_rec prev_rels (c,cty) = let (cty,direct_rels) = minimal_free_rels env sigma (c,cty) in let combined_rels = Int.Set.union prev_rels direct_rels in - let folder rels i = snd (minimalrec_free_rels_rec rels (c, EConstr.of_constr (unsafe_type_of env sigma (mkRel i)))) + let folder rels i = snd (minimalrec_free_rels_rec rels (c, unsafe_type_of env sigma (mkRel i))) in (cty, List.fold_left folder combined_rels (Int.Set.elements (Int.Set.diff direct_rels prev_rels))) in minimalrec_free_rels_rec Int.Set.empty @@ -1184,7 +1182,7 @@ let sig_clausal_form env sigma sort_of_ty siglen ty dflt = (* is the default value typable with the expected type *) let dflt_typ = unsafe_type_of env sigma dflt in try - let () = evdref := Evarconv.the_conv_x_leq env (EConstr.of_constr dflt_typ) p_i !evdref in + let () = evdref := Evarconv.the_conv_x_leq env dflt_typ p_i !evdref in let () = evdref := Evarconv.consider_remaining_unif_problems env !evdref in dflt with Evarconv.UnableToUnify _ -> @@ -1200,7 +1198,6 @@ let sig_clausal_form env sigma sort_of_ty siglen ty dflt = match evopt with | Some w -> let w_type = unsafe_type_of env !evdref w in - let w_type = EConstr.of_constr w_type in if Evarconv.e_cumul env evdref w_type a then let exist_term = Evarutil.evd_comb1 (Evd.fresh_global env) evdref sigdata.intro in let exist_term = EConstr.of_constr exist_term in @@ -1290,7 +1287,7 @@ let make_iterated_tuple env sigma dflt (z,zty) = sigma, (tuple,tuplety,dfltval) let rec build_injrec env sigma dflt c = function - | [] -> make_iterated_tuple env sigma dflt (c,EConstr.of_constr (unsafe_type_of env sigma c)) + | [] -> make_iterated_tuple env sigma dflt (c,unsafe_type_of env sigma c) | ((sp,cnum),argnum)::l -> try let (cnum_nlams,cnum_env,kont) = descend_then env sigma c cnum in @@ -1345,7 +1342,7 @@ let inject_if_homogenous_dependent_pair ty = if not (Ind_tables.check_scheme (!eq_dec_scheme_kind_name()) (fst ind) && pf_apply is_conv gl ar1.(2) ar2.(2)) then raise Exit; Coqlib.check_required_library ["Coq";"Logic";"Eqdep_dec"]; - let new_eq_args = [|EConstr.of_constr (pf_unsafe_type_of gl ar1.(3));ar1.(3);ar2.(3)|] in + let new_eq_args = [|pf_unsafe_type_of gl ar1.(3);ar1.(3);ar2.(3)|] in let inj2 = Coqlib.coq_constant "inj_pair2_eq_dec is missing" ["Logic";"Eqdep_dec"] "inj_pair2_eq_dec" in let inj2 = EConstr.of_constr inj2 in @@ -1613,7 +1610,6 @@ let cutSubstInHyp l2r eqn id = let sigma = Proofview.Goal.sigma gl in let (lbeq,u,(t,e1,e2)) = find_eq_data_decompose gl eqn in let typ = pf_get_hyp_typ id gl in - let typ = EConstr.of_constr typ in let (e1,e2) = if l2r then (e1,e2) else (e2,e1) in let Sigma ((typ, expected), sigma, p) = subst_tuple_term env sigma e1 e2 typ in let tac = @@ -1702,8 +1698,7 @@ let is_eq_x gl x d = | Var id' -> Id.equal id id' | _ -> false in - let c = pf_nf_evar gl (NamedDecl.get_type d) in - let c = EConstr.of_constr c in + let c = pf_nf_evar gl (EConstr.of_constr (NamedDecl.get_type d)) in let (_,lhs,rhs) = pi3 (find_eq_data_decompose gl c) in if (is_var x lhs) && not (local_occur_var (project gl) x rhs) then raise (FoundHyp (id,rhs,true)); if (is_var x rhs) && not (local_occur_var (project gl) x lhs) then raise (FoundHyp (id,lhs,false)) @@ -1852,7 +1847,6 @@ let subst_all ?(flags=default_subst_tactic_flags ()) () = let find_eq_data_decompose = find_eq_data_decompose gl in let test (_,c) = try - let c = EConstr.of_constr c in let lbeq,u,(_,x,y) = find_eq_data_decompose c in let eq = Universes.constr_of_global_univ (lbeq.eq,u) in if flags.only_leibniz then restrict_to_eq_and_identity eq; diff --git a/tactics/hints.ml b/tactics/hints.ml index 2446b6996..851e9f01f 100644 --- a/tactics/hints.ml +++ b/tactics/hints.ml @@ -919,7 +919,7 @@ let make_mode ref m = let make_trivial env sigma poly ?(name=PathAny) r = let c,ctx = fresh_global_or_constr env sigma poly r in let sigma = Evd.merge_context_set univ_flexible sigma ctx in - let t = hnf_constr env sigma (EConstr.of_constr (unsafe_type_of env sigma c)) in + let t = hnf_constr env sigma (unsafe_type_of env sigma c) in let hd = head_constr sigma t in let ce = mk_clenv_from_env env sigma None (c,t) in (Some hd, { pri=1; @@ -1239,7 +1239,6 @@ let interp_hints poly = let sigma = Evd.from_env env in let f poly c = let evd,c = Constrintern.interp_open_constr env sigma c in - let c = EConstr.of_constr c in prepare_hint true (poly,false) (Global.env()) Evd.empty (evd,c) in let fref r = let gr = global_with_alias r in diff --git a/tactics/hipattern.ml b/tactics/hipattern.ml index fa114a3d3..607d6d2a9 100644 --- a/tactics/hipattern.ml +++ b/tactics/hipattern.ml @@ -442,7 +442,7 @@ let find_eq_data sigma eqn = (* fails with PatternMatchingFailure *) let extract_eq_args gl = function | MonomorphicLeibnizEq (e1,e2) -> - let t = pf_unsafe_type_of gl e1 in (EConstr.of_constr t,e1,e2) + let t = pf_unsafe_type_of gl e1 in (t,e1,e2) | PolymorphicLeibnizEq (t,e1,e2) -> (t,e1,e2) | HeterogenousEq (t1,e1,t2,e2) -> if pf_conv_x gl t1 t2 then (t1,e1,e2) diff --git a/tactics/inv.ml b/tactics/inv.ml index a398e04dd..426749a75 100644 --- a/tactics/inv.ml +++ b/tactics/inv.ml @@ -346,7 +346,7 @@ let projectAndApply as_mode thin avoid id eqname names depids = let sigma = project gl in (** We only look at the type of hypothesis "id" *) let hyp = pf_nf_evar gl (pf_get_hyp_typ id (Proofview.Goal.assume gl)) in - let (t,t1,t2) = Hipattern.dest_nf_eq gl (EConstr.of_constr hyp) in + let (t,t1,t2) = Hipattern.dest_nf_eq gl hyp in match (EConstr.kind sigma t1, EConstr.kind sigma t2) with | Var id1, _ -> generalizeRewriteIntros as_mode (subst_hyp true id) depids id1 | _, Var id2 -> generalizeRewriteIntros as_mode (subst_hyp false id) depids id2 @@ -443,7 +443,7 @@ let raw_inversion inv_kind id status names = let concl = Proofview.Goal.concl gl in let c = mkVar id in let (ind, t) = - try pf_apply Tacred.reduce_to_atomic_ind gl (EConstr.of_constr (pf_unsafe_type_of gl c)) + try pf_apply Tacred.reduce_to_atomic_ind gl (pf_unsafe_type_of gl c) with UserError _ -> let msg = str "The type of " ++ pr_id id ++ str " is not inductive." in CErrors.user_err msg diff --git a/tactics/leminv.ml b/tactics/leminv.ml index 2d59285e6..3199623e7 100644 --- a/tactics/leminv.ml +++ b/tactics/leminv.ml @@ -258,7 +258,6 @@ let add_inversion_lemma_exn na com comsort bool tac = let env = Global.env () in let evd = ref (Evd.from_env env) in let c = Constrintern.interp_type_evars env evd com in - let c = EConstr.of_constr c in let sigma, sort = Pretyping.interp_sort !evd comsort in try add_inversion_lemma na env sigma c sort bool tac diff --git a/tactics/tacticals.ml b/tactics/tacticals.ml index d79a74b36..89acc149c 100644 --- a/tactics/tacticals.ml +++ b/tactics/tacticals.ml @@ -128,7 +128,7 @@ let onClauseLR tac cl gls = tclMAP tac (List.rev (Locusops.simple_clause_of hyps cl)) gls let ifOnHyp pred tac1 tac2 id gl = - if pred (id,EConstr.of_constr (pf_get_hyp_typ gl id)) then + if pred (id,pf_get_hyp_typ gl id) then tac1 id gl else tac2 id gl @@ -248,10 +248,10 @@ let compute_constructor_signatures isrec ((_,k as ity),u) = Array.map2 analrec lc lrecargs let elimination_sort_of_goal gl = - pf_apply Retyping.get_sort_family_of gl (EConstr.of_constr (pf_concl gl)) + pf_apply Retyping.get_sort_family_of gl (pf_concl gl) let elimination_sort_of_hyp id gl = - pf_apply Retyping.get_sort_family_of gl (EConstr.of_constr (pf_get_hyp_typ gl id)) + pf_apply Retyping.get_sort_family_of gl (pf_get_hyp_typ gl id) let elimination_sort_of_clause = function | None -> elimination_sort_of_goal @@ -269,21 +269,22 @@ let pf_constr_of_global gr k = let gl_make_elim ind gl = let gr = Indrec.lookup_eliminator (fst ind) (elimination_sort_of_goal gl) in - pf_apply Evd.fresh_global gl gr + let (sigma, c) = pf_apply Evd.fresh_global gl gr in + (sigma, EConstr.of_constr c) let gl_make_case_dep ind gl = let sigma = Sigma.Unsafe.of_evar_map (Tacmach.project gl) in let Sigma (r, sigma, _) = Indrec.build_case_analysis_scheme (pf_env gl) sigma ind true (elimination_sort_of_goal gl) in - (Sigma.to_evar_map sigma, r) + (Sigma.to_evar_map sigma, EConstr.of_constr r) let gl_make_case_nodep ind gl = let sigma = Sigma.Unsafe.of_evar_map (Tacmach.project gl) in let Sigma (r, sigma, _) = Indrec.build_case_analysis_scheme (pf_env gl) sigma ind false (elimination_sort_of_goal gl) in - (Sigma.to_evar_map sigma, r) + (Sigma.to_evar_map sigma, EConstr.of_constr r) let make_elim_branch_assumptions ba gl = let assums = @@ -583,7 +584,6 @@ module New = struct let ifOnHyp pred tac1 tac2 id = Proofview.Goal.nf_enter { enter = begin fun gl -> let typ = Tacmach.New.pf_get_hyp_typ id gl in - let typ = EConstr.of_constr typ in if pred (id,typ) then tac1 id else @@ -630,7 +630,7 @@ module New = struct (Proofview.Goal.nf_enter { enter = begin fun gl -> let indclause = Tacmach.New.of_old (fun gl -> mk_clenv_from gl (c, t)) gl in (* applying elimination_scheme just a little modified *) - let elimclause = Tacmach.New.of_old (fun gls -> mk_clenv_from gls (EConstr.of_constr elim,EConstr.of_constr (Tacmach.New.pf_unsafe_type_of gl (EConstr.of_constr elim)))) gl in + let elimclause = Tacmach.New.of_old (fun gls -> mk_clenv_from gls (elim,Tacmach.New.pf_unsafe_type_of gl elim)) gl in let indmv = match EConstr.kind elimclause.evd (last_arg elimclause.evd elimclause.templval.Evd.rebus) with | Meta mv -> mv @@ -642,7 +642,7 @@ module New = struct | Meta p -> p | _ -> let name_elim = - match kind_of_term elim with + match EConstr.kind sigma elim with | Const (kn, _) -> string_of_con kn | Var id -> string_of_id id | _ -> "\b" @@ -680,7 +680,7 @@ module New = struct let elimination_then tac c = Proofview.Goal.nf_enter { enter = begin fun gl -> - let (ind,t) = pf_reduce_to_quantified_ind gl (EConstr.of_constr (pf_unsafe_type_of gl c)) in + let (ind,t) = pf_reduce_to_quantified_ind gl (pf_unsafe_type_of gl c) in let isrec,mkelim = match (Global.lookup_mind (fst (fst ind))).mind_record with | None -> true,gl_make_elim @@ -715,7 +715,7 @@ module New = struct let elimination_sort_of_hyp id gl = (** Retyping will expand evars anyway. *) let c = pf_get_hyp_typ id (Goal.assume gl) in - pf_apply Retyping.get_sort_family_of gl (EConstr.of_constr c) + pf_apply Retyping.get_sort_family_of gl c let elimination_sort_of_clause id gl = match id with | None -> elimination_sort_of_goal gl diff --git a/tactics/tactics.ml b/tactics/tactics.ml index dcaa15fd8..f79f7f1a8 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -1260,7 +1260,6 @@ let cut c = try (** Backward compat: ensure that [c] is well-typed. *) let typ = Typing.unsafe_type_of env sigma c in - let typ = EConstr.of_constr typ in let typ = whd_all env sigma typ in match EConstr.kind sigma typ with | Sort _ -> true @@ -1515,7 +1514,7 @@ let find_ind_eliminator ind s gl = evd, c let find_eliminator c gl = - let ((ind,u),t) = Tacmach.New.pf_reduce_to_quantified_ind gl (EConstr.of_constr (Tacmach.New.pf_unsafe_type_of gl c)) in + let ((ind,u),t) = Tacmach.New.pf_reduce_to_quantified_ind gl (Tacmach.New.pf_unsafe_type_of gl c) in if is_nonrec ind then raise IsNonrec; let evd, c = find_ind_eliminator ind (Tacticals.New.elimination_sort_of_goal gl) gl in evd, {elimindex = None; elimbody = (c,NoBindings); @@ -1891,7 +1890,6 @@ let apply_in_once sidecond_first with_delta with_destruct with_evars naming let flags = if with_delta then default_unify_flags () else default_no_delta_unify_flags () in let t' = Tacmach.New.pf_get_hyp_typ id gl in - let t' = EConstr.of_constr t' in let innerclause = mk_clenv_from_env env sigma (Some 0) (mkVar id,t') in let targetid = find_name true (local_assum (Anonymous,t')) naming gl in let rec aux idstoclear with_destruct c = @@ -1949,7 +1947,7 @@ let apply_in_delayed_once sidecond_first with_delta with_destruct with_evars nam let cut_and_apply c = Proofview.Goal.nf_enter { enter = begin fun gl -> let sigma = Tacmach.New.project gl in - match EConstr.kind sigma (Tacmach.New.pf_hnf_constr gl (EConstr.of_constr (Tacmach.New.pf_unsafe_type_of gl c))) with + match EConstr.kind sigma (Tacmach.New.pf_hnf_constr gl (Tacmach.New.pf_unsafe_type_of gl c)) with | Prod (_,c1,c2) when Vars.noccurn sigma 1 c2 -> let concl = Proofview.Goal.concl gl in let env = Tacmach.New.pf_env gl in @@ -2004,7 +2002,7 @@ let exact_proof c = Proofview.Goal.nf_enter { enter = begin fun gl -> Refine.refine { run = begin fun sigma -> let sigma = Sigma.to_evar_map sigma in - let (c, ctx) = Constrintern.interp_casted_constr (pf_env gl) sigma c (EConstr.Unsafe.to_constr (pf_concl gl)) in + let (c, ctx) = Constrintern.interp_casted_constr (pf_env gl) sigma c (pf_concl gl) in let c = EConstr.of_constr c in let sigma = Evd.merge_universe_context sigma ctx in Sigma.Unsafe.of_pair (c, sigma) @@ -2326,7 +2324,6 @@ let intro_decomp_eq loc l thin tac id = Proofview.Goal.nf_enter { enter = begin fun gl -> let c = mkVar id in let t = Tacmach.New.pf_unsafe_type_of gl c in - let t = EConstr.of_constr t in let _,t = Tacmach.New.pf_reduce_to_quantified_ind gl t in match my_find_eq_data_decompose gl t with | Some (eq,u,eq_args) -> @@ -2341,7 +2338,6 @@ let intro_or_and_pattern loc with_evars bracketed ll thin tac id = Proofview.Goal.enter { enter = begin fun gl -> let c = mkVar id in let t = Tacmach.New.pf_unsafe_type_of gl c in - let t = EConstr.of_constr t in let (ind,t) = Tacmach.New.pf_reduce_to_quantified_ind gl t in let branchsigns = compute_constructor_signatures false ind in let nv_with_let = Array.map List.length branchsigns in @@ -2366,7 +2362,7 @@ let rewrite_hyp_then assert_style with_evars thin l2r id tac = let sigma = Tacmach.New.project gl in let type_of = Tacmach.New.pf_unsafe_type_of gl in let whd_all = Tacmach.New.pf_apply whd_all gl in - let t = whd_all (EConstr.of_constr (type_of (mkVar id))) in + let t = whd_all (type_of (mkVar id)) in let eqtac, thin = match match_with_equality_type sigma t with | Some (hdcncl,[_;lhs;rhs]) -> if l2r && isVar sigma lhs && not (occur_var env sigma (destVar sigma lhs) rhs) then @@ -2774,7 +2770,6 @@ let forward b usetac ipat c = | None -> Proofview.Goal.enter { enter = begin fun gl -> let t = Tacmach.New.pf_unsafe_type_of gl c in - let t = EConstr.of_constr t in let sigma = Tacmach.New.project gl in let hd = head_ident sigma c in Tacticals.New.tclTHENFIRST (assert_as true hd ipat t) (exact_no_check c) @@ -2861,7 +2856,7 @@ let old_generalize_dep ?(with_let=false) c gl = -> id::tothin | _ -> tothin in - let cl' = it_mkNamedProd_or_LetIn (EConstr.of_constr (Tacmach.pf_concl gl)) to_quantify in + let cl' = it_mkNamedProd_or_LetIn (Tacmach.pf_concl gl) to_quantify in let body = if with_let then match EConstr.kind sigma c with @@ -3222,7 +3217,6 @@ let atomize_param_of_ind_then (indref,nparams,_) hyp0 tac = let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in let tmptyp0 = Tacmach.New.pf_get_hyp_typ hyp0 (Proofview.Goal.assume gl) in - let tmptyp0 = EConstr.of_constr tmptyp0 in let reduce_to_quantified_ref = Tacmach.New.pf_apply reduce_to_quantified_ref gl in let typ0 = reduce_to_quantified_ref indref tmptyp0 in let prods, indtyp = decompose_prod_assum sigma typ0 in @@ -3266,7 +3260,7 @@ let atomize_param_of_ind_then (indref,nparams,_) hyp0 tac = | Var id -> id | _ -> let type_of = Tacmach.New.pf_unsafe_type_of gl in - id_of_name_using_hdchar (Global.env()) (type_of c) Anonymous in + id_of_name_using_hdchar (Global.env()) (EConstr.Unsafe.to_constr (type_of c)) Anonymous in let x = fresh_id_in_env avoid id env in Tacticals.New.tclTHEN (letin_tac None (Name x) c None allHypsAndConcl) @@ -3660,7 +3654,6 @@ let abstract_args gl generalize_vars dep id defined f args = let sigma = ref (Tacmach.project gl) in let env = Tacmach.pf_env gl in let concl = Tacmach.pf_concl gl in - let concl = EConstr.of_constr concl in let dep = dep || local_occur_var !sigma id concl in let avoid = ref [] in let get_id name = @@ -3681,7 +3674,6 @@ let abstract_args gl generalize_vars dep id defined f args = in let ty = EConstr.of_constr ty in let argty = Tacmach.pf_unsafe_type_of gl arg in - let argty = EConstr.of_constr argty in let sigma', ty = Evarsolve.refresh_universes (Some true) env !sigma ty in let () = sigma := sigma' in let lenctx = List.length ctx in @@ -3723,7 +3715,6 @@ let abstract_args gl generalize_vars dep id defined f args = in if dogen then let tyf' = Tacmach.pf_unsafe_type_of gl f' in - let tyf' = EConstr.of_constr tyf' in let arity, ctx, ctxenv, c', args, eqs, refls, nogen, vars, env = Array.fold_left aux (tyf',[],env,f',[],[],[],Id.Set.empty,Id.Set.empty,env) args' in @@ -3739,7 +3730,6 @@ let abstract_args gl generalize_vars dep id defined f args = else None, c' in let typ = Tacmach.pf_get_hyp_typ gl id in - let typ = EConstr.of_constr typ in let tac = make_abstract_generalize (pf_env gl) id typ concl dep ctx body c' eqs args refls in let tac = Proofview.Unsafe.tclEVARS !sigma <*> tac in Some (tac, dep, succ (List.length ctx), vars) @@ -3797,7 +3787,6 @@ let specialize_eqs id gl = let open Context.Rel.Declaration in let env = Tacmach.pf_env gl in let ty = Tacmach.pf_get_hyp_typ gl id in - let ty = EConstr.of_constr ty in let evars = ref (project gl) in let unif env evars c1 c2 = compare_upto_variables !evars c1 c2 && Evarconv.e_conv env evars c1 c2 @@ -4062,7 +4051,6 @@ let compute_elim_signature (evd,(elimc,elimt),ind_type_guess) names_info = let guess_elim isrec dep s hyp0 gl = let tmptyp0 = Tacmach.New.pf_get_hyp_typ hyp0 gl in - let tmptyp0 = EConstr.of_constr tmptyp0 in let mind,_ = Tacmach.New.pf_reduce_to_quantified_ind gl tmptyp0 in let evd, elimc = if isrec && not (is_nonrec (fst mind)) then find_ind_eliminator (fst mind) s gl @@ -4080,16 +4068,13 @@ let guess_elim isrec dep s hyp0 gl = (Sigma.to_evar_map sigma, ind) in let elimt = Tacmach.New.pf_unsafe_type_of gl elimc in - let elimt = EConstr.of_constr elimt in evd, ((elimc, NoBindings), elimt), mkIndU mind let given_elim hyp0 (elimc,lbind as e) gl = let sigma = Tacmach.New.project gl in let tmptyp0 = Tacmach.New.pf_get_hyp_typ hyp0 gl in - let tmptyp0 = EConstr.of_constr tmptyp0 in let ind_type_guess,_ = decompose_app sigma (snd (decompose_prod sigma tmptyp0)) in let elimt = Tacmach.New.pf_unsafe_type_of gl elimc in - let elimt = EConstr.of_constr elimt in Tacmach.New.project gl, (e, elimt), ind_type_guess type scheme_signature = @@ -4127,7 +4112,7 @@ let get_elim_signature elim hyp0 gl = let is_functional_induction elimc gl = let sigma = Tacmach.New.project gl in - let scheme = compute_elim_sig sigma ~elimc (EConstr.of_constr (Tacmach.New.pf_unsafe_type_of gl (fst elimc))) in + let scheme = compute_elim_sig sigma ~elimc (Tacmach.New.pf_unsafe_type_of gl (fst elimc)) in (* The test is not safe: with non-functional induction on non-standard induction scheme, this may fail *) Option.is_empty scheme.indarg @@ -4162,7 +4147,6 @@ let recolle_clenv i params args elimclause gl = arr in let k = match i with -1 -> Array.length lindmv - List.length args | _ -> i in (* parameters correspond to first elts of lid. *) - let pf_get_hyp_typ id gl = EConstr.of_constr (pf_get_hyp_typ id gl) in let clauses_params = List.map_i (fun i id -> mkVar id , pf_get_hyp_typ id gl, lindmv.(i)) 0 params in @@ -4523,7 +4507,7 @@ let induction_gen_l isrec with_evars elim names lc = let type_of = Tacmach.New.pf_unsafe_type_of gl in let sigma = Tacmach.New.project gl in let x = - id_of_name_using_hdchar (Global.env()) (type_of c) Anonymous in + id_of_name_using_hdchar (Global.env()) (EConstr.Unsafe.to_constr (type_of c)) Anonymous in let id = new_fresh_id [] x gl in let newl' = List.map (fun r -> replace_term sigma c (mkVar id) r) l' in @@ -4778,7 +4762,6 @@ let symmetry_in id = Proofview.Goal.enter { enter = begin fun gl -> let sigma = Tacmach.New.project gl in let ctype = Tacmach.New.pf_unsafe_type_of gl (mkVar id) in - let ctype = EConstr.of_constr ctype in let sign,t = decompose_prod_assum sigma ctype in Proofview.tclORELSE begin @@ -4832,7 +4815,6 @@ let prove_transitivity hdcncl eq_kind t = let sigma = Tacmach.New.project gl in let type_of = Typing.unsafe_type_of env sigma in let typt = type_of t in - let typt = EConstr.of_constr typt in (mkApp(hdcncl, [| typ1; c1; typt ;t |]), mkApp(hdcncl, [| typt; t; typ2; c2 |])) in diff --git a/toplevel/auto_ind_decl.ml b/toplevel/auto_ind_decl.ml index 5f33ae834..5e686c41e 100644 --- a/toplevel/auto_ind_decl.ml +++ b/toplevel/auto_ind_decl.ml @@ -321,11 +321,12 @@ let _ = beq_scheme_kind_aux := fun () -> beq_scheme_kind (* This function tryies to get the [inductive] between a constr the constr should be Ind i or App(Ind i,[|args|]) *) -let destruct_ind c = - try let u,v = destApp c in - let indc = destInd u in +let destruct_ind sigma c = + let open EConstr in + try let u,v = destApp sigma c in + let indc = destInd sigma u in indc,v - with DestKO -> let indc = destInd c in + with DestKO -> let indc = destInd sigma c in indc,[||] (* @@ -338,11 +339,12 @@ so from Ai we can find the the correct eq_Ai bl_ai or lb_ai *) (* used in the leib -> bool side*) let do_replace_lb mode lb_scheme_key aavoid narg p q = + let open EConstr in let avoid = Array.of_list aavoid in - let do_arg v offset = + let do_arg sigma v offset = try let x = narg*offset in - let s = destVar v in + let s = destVar sigma v in let n = Array.length avoid in let rec find i = if Id.equal avoid.(n-i) s then avoid.(n-i-x) @@ -355,7 +357,7 @@ let do_replace_lb mode lb_scheme_key aavoid narg p q = (* if this happen then the args have to be already declared as a Parameter*) ( - let mp,dir,lbl = repr_con (fst (destConst v)) in + let mp,dir,lbl = repr_con (fst (destConst sigma v)) in mkConst (make_con mp dir (mk_label ( if Int.equal offset 1 then ("eq_"^(Label.to_string lbl)) else ((Label.to_string lbl)^"_lb") @@ -364,7 +366,8 @@ let do_replace_lb mode lb_scheme_key aavoid narg p q = in Proofview.Goal.nf_enter { enter = begin fun gl -> let type_of_pq = Tacmach.New.of_old (fun gl -> pf_unsafe_type_of gl p) gl in - let u,v = destruct_ind type_of_pq + let sigma = Tacmach.New.project gl in + let u,v = destruct_ind sigma type_of_pq in let lb_type_of_p = try let c, eff = find_scheme ~mode lb_scheme_key (out_punivs u) (*FIXME*) in @@ -376,20 +379,20 @@ let do_replace_lb mode lb_scheme_key aavoid narg p q = (str "Leibniz->boolean:" ++ str "You have to declare the" ++ str "decidability over " ++ - Printer.pr_constr type_of_pq ++ + Printer.pr_econstr type_of_pq ++ str " first.") in Tacticals.New.tclZEROMSG err_msg in lb_type_of_p >>= fun (lb_type_of_p,eff) -> + Proofview.tclEVARMAP >>= fun sigma -> let lb_args = Array.append (Array.append (Array.map (fun x -> x) v) - (Array.map (fun x -> do_arg x 1) v)) - (Array.map (fun x -> do_arg x 2) v) - in let app = if Array.equal Term.eq_constr lb_args [||] + (Array.map (fun x -> do_arg sigma x 1) v)) + (Array.map (fun x -> do_arg sigma x 2) v) + in let app = if Array.is_empty lb_args then lb_type_of_p else mkApp (lb_type_of_p,lb_args) in - let app = EConstr.of_constr app in Tacticals.New.tclTHENLIST [ Proofview.tclEFFECTS eff; Equality.replace p q ; apply app ; Auto.default_auto] @@ -397,13 +400,12 @@ let do_replace_lb mode lb_scheme_key aavoid narg p q = (* used in the bool -> leib side *) let do_replace_bl mode bl_scheme_key (ind,u as indu) aavoid narg lft rgt = - let lft = EConstr.Unsafe.to_constr lft in - let rgt = EConstr.Unsafe.to_constr rgt in + let open EConstr in let avoid = Array.of_list aavoid in - let do_arg v offset = + let do_arg sigma v offset = try let x = narg*offset in - let s = destVar v in + let s = destVar sigma v in let n = Array.length avoid in let rec find i = if Id.equal avoid.(n-i) s then avoid.(n-i-x) @@ -416,7 +418,7 @@ let do_replace_bl mode bl_scheme_key (ind,u as indu) aavoid narg lft rgt = (* if this happen then the args have to be already declared as a Parameter*) ( - let mp,dir,lbl = repr_con (fst (destConst v)) in + let mp,dir,lbl = repr_con (fst (destConst sigma v)) in mkConst (make_con mp dir (mk_label ( if Int.equal offset 1 then ("eq_"^(Label.to_string lbl)) else ((Label.to_string lbl)^"_bl") @@ -427,13 +429,12 @@ let do_replace_bl mode bl_scheme_key (ind,u as indu) aavoid narg lft rgt = let rec aux l1 l2 = match (l1,l2) with | (t1::q1,t2::q2) -> - let t1 = EConstr.of_constr t1 in - let t2 = EConstr.of_constr t2 in Proofview.Goal.enter { enter = begin fun gl -> let tt1 = Tacmach.New.pf_unsafe_type_of gl t1 in - if EConstr.eq_constr (Tacmach.New.project gl) t1 t2 then aux q1 q2 + let sigma = Tacmach.New.project gl in + if EConstr.eq_constr sigma t1 t2 then aux q1 q2 else ( - let u,v = try destruct_ind tt1 + let u,v = try destruct_ind sigma tt1 (* trick so that the good sequence is returned*) with e when CErrors.noncritical e -> indu,[||] in if eq_ind (fst u) ind @@ -450,20 +451,19 @@ let do_replace_bl mode bl_scheme_key (ind,u as indu) aavoid narg lft rgt = (str "boolean->Leibniz:" ++ str "You have to declare the" ++ str "decidability over " ++ - Printer.pr_constr tt1 ++ + Printer.pr_econstr tt1 ++ str " first.") in error err_msg in let bl_args = Array.append (Array.append (Array.map (fun x -> x) v) - (Array.map (fun x -> do_arg x 1) v)) - (Array.map (fun x -> do_arg x 2) v ) + (Array.map (fun x -> do_arg sigma x 1) v)) + (Array.map (fun x -> do_arg sigma x 2) v ) in - let app = if Array.equal Term.eq_constr bl_args [||] + let app = if Array.is_empty bl_args then bl_t1 else mkApp (bl_t1,bl_args) in - let app = EConstr.of_constr app in Tacticals.New.tclTHENLIST [ Proofview.tclEFFECTS eff; Equality.replace_by t1 t2 @@ -475,21 +475,22 @@ let do_replace_bl mode bl_scheme_key (ind,u as indu) aavoid narg lft rgt = | ([],[]) -> Proofview.tclUNIT () | _ -> Tacticals.New.tclZEROMSG (str "Both side of the equality must have the same arity.") in - begin try Proofview.tclUNIT (destApp lft) + Proofview.tclEVARMAP >>= fun sigma -> + begin try Proofview.tclUNIT (destApp sigma lft) with DestKO -> Tacticals.New.tclZEROMSG (str "replace failed.") end >>= fun (ind1,ca1) -> - begin try Proofview.tclUNIT (destApp rgt) + begin try Proofview.tclUNIT (destApp sigma rgt) with DestKO -> Tacticals.New.tclZEROMSG (str "replace failed.") end >>= fun (ind2,ca2) -> - begin try Proofview.tclUNIT (out_punivs (destInd ind1)) + begin try Proofview.tclUNIT (out_punivs (destInd sigma ind1)) with DestKO -> - begin try Proofview.tclUNIT (fst (fst (destConstruct ind1))) + begin try Proofview.tclUNIT (fst (fst (destConstruct sigma ind1))) with DestKO -> Tacticals.New.tclZEROMSG (str "The expected type is an inductive one.") end end >>= fun (sp1,i1) -> - begin try Proofview.tclUNIT (out_punivs (destInd ind2)) + begin try Proofview.tclUNIT (out_punivs (destInd sigma ind2)) with DestKO -> - begin try Proofview.tclUNIT (fst (fst (destConstruct ind2))) + begin try Proofview.tclUNIT (fst (fst (destConstruct sigma ind2))) with DestKO -> Tacticals.New.tclZEROMSG (str "The expected type is an inductive one.") end end >>= fun (sp2,i2) -> diff --git a/toplevel/class.ml b/toplevel/class.ml index e55489471..4c7aa01cd 100644 --- a/toplevel/class.ml +++ b/toplevel/class.ml @@ -207,7 +207,7 @@ let build_id_coercion idf_opt source poly = let _ = if not (Reductionops.is_conv_leq env sigma - (EConstr.of_constr (Typing.unsafe_type_of env sigma (EConstr.of_constr val_f))) (EConstr.of_constr typ_f)) + (Typing.unsafe_type_of env sigma (EConstr.of_constr val_f)) (EConstr.of_constr typ_f)) then user_err (strbrk "Cannot be defined as coercion (maybe a bad number of arguments).") diff --git a/toplevel/classes.ml b/toplevel/classes.ml index acfbce357..5087aa0c8 100644 --- a/toplevel/classes.ml +++ b/toplevel/classes.ml @@ -83,7 +83,7 @@ let type_ctx_instance evars env ctx inst subst = let t' = substl subst (RelDecl.get_type decl) in let c', l = match decl with - | LocalAssum _ -> interp_casted_constr_evars env evars (List.hd l) t', List.tl l + | LocalAssum _ -> EConstr.Unsafe.to_constr (interp_casted_constr_evars env evars (List.hd l) t'), List.tl l | LocalDef (_,b,_) -> substl subst b, l in let d = RelDecl.get_name decl, Some c', t' in @@ -151,6 +151,7 @@ let new_instance ?(abstract=false) ?(global=false) ?(refine= !refine_instance) p let k, u, cty, ctx', ctx, len, imps, subst = let impls, ((env', ctx), imps) = interp_context_evars env evars ctx in let c', imps' = interp_type_evars_impls ~impls env' evars tclass in + let c' = EConstr.Unsafe.to_constr c' in let len = List.length ctx in let imps = imps @ Impargs.lift_implicits len imps' in let ctx', c = decompose_prod_assum c' in @@ -219,6 +220,7 @@ let new_instance ?(abstract=false) ?(global=false) ?(refine= !refine_instance) p | None -> if List.is_empty k.cl_props then Some (Inl subst) else None | Some (Inr term) -> let c = interp_casted_constr_evars env' evars term cty in + let c = EConstr.Unsafe.to_constr c in Some (Inr (c, subst)) | Some (Inl props) -> let get_id = diff --git a/toplevel/command.ml b/toplevel/command.ml index 1fc89b8b0..7e3828131 100644 --- a/toplevel/command.ml +++ b/toplevel/command.ml @@ -103,6 +103,7 @@ let interp_definition pl bl p red_option c ctypopt = let ctx = Context.Rel.map (Vars.subst_univs_constr subst) ctx in let env_bl = push_rel_context ctx env in let c, imps2 = interp_constr_evars_impls ~impls env_bl evdref c in + let c = EConstr.Unsafe.to_constr c in let nf,subst = Evarutil.e_nf_evars_and_universes evdref in let body = nf (it_mkLambda_or_LetIn c ctx) in let vars = Universes.universes_of_constr body in @@ -116,8 +117,10 @@ let interp_definition pl bl p red_option c ctypopt = let ctx = Context.Rel.map (Vars.subst_univs_constr subst) ctx in let env_bl = push_rel_context ctx env in let c, imps2 = interp_casted_constr_evars_impls ~impls env_bl evdref c ty in + let c = EConstr.Unsafe.to_constr c in let nf, subst = Evarutil.e_nf_evars_and_universes evdref in let body = nf (it_mkLambda_or_LetIn c ctx) in + let ty = EConstr.Unsafe.to_constr ty in let typ = nf (Term.it_mkProd_or_LetIn ty ctx) in let beq b1 b2 = if b1 then b2 else not b2 in let impl_eq (x,y,z) (x',y',z') = beq x x' && beq y y' && beq z z' in @@ -264,6 +267,7 @@ match local with let interp_assumption evdref env impls bl c = let c = prod_constr_expr c bl in let ty, impls = interp_type_evars_impls env evdref ~impls c in + let ty = EConstr.Unsafe.to_constr ty in let evd, nf = nf_evars_and_universes !evdref in let ctx = Evd.universe_context_set evd in ((nf ty, ctx), impls) @@ -318,6 +322,7 @@ let do_assumptions_bound_univs coe kind nl id pl c = let evdref = ref (Evd.from_ctx ctx) in let ty, impls = interp_type_evars_impls env evdref c in let nf, subst = Evarutil.e_nf_evars_and_universes evdref in + let ty = EConstr.Unsafe.to_constr ty in let ty = nf ty in let vars = Universes.universes_of_constr ty in let evd = Evd.restrict_universe_context !evdref vars in @@ -439,9 +444,10 @@ let interp_ind_arity env evdref ind = let imps = Implicit_quantifiers.implicits_of_glob_constr ~with_products:true c in let t, impls = understand_tcc_evars env evdref ~expected_type:IsType c, imps in let pseudo_poly = check_anonymous_type c in - let () = if not (Reduction.is_arity env t) then + let () = if not (Reductionops.is_arity env !evdref t) then user_err ~loc:(constr_loc ind.ind_arity) (str "Not an arity") in + let t = EConstr.Unsafe.to_constr t in t, pseudo_poly, impls let interp_cstrs evdref env impls mldata arity ind = @@ -449,7 +455,7 @@ let interp_cstrs evdref env impls mldata arity ind = (* Complete conclusions of constructor types if given in ML-style syntax *) let ctyps' = List.map2 (complete_conclusion mldata) cnames ctyps in (* Interpret the constructor types *) - let ctyps'', cimpls = List.split (List.map (interp_type_evars_impls evdref env ~impls) ctyps') in + let ctyps'', cimpls = List.split (List.map (interp_type_evars_impls evdref env ~impls %> on_fst EConstr.Unsafe.to_constr) ctyps') in (cnames, ctyps'', cimpls) let sign_level env evd sign = @@ -842,12 +848,14 @@ let interp_fix_context env evdref isfix fix = ((env'', ctx' @ ctx), (impl_env',imps @ imps'), annot) let interp_fix_ccl evdref impls (env,_) fix = - interp_type_evars_impls ~impls env evdref fix.fix_type + let (c, impl) = interp_type_evars_impls ~impls env evdref fix.fix_type in + (EConstr.Unsafe.to_constr c, impl) let interp_fix_body env_rec evdref impls (_,ctx) fix ccl = Option.map (fun body -> let env = push_rel_context ctx env_rec in let body = interp_casted_constr_evars env evdref ~impls body ccl in + let body = EConstr.Unsafe.to_constr body in it_mkLambda_or_LetIn body ctx) fix.fix_body let build_fix_type (_,ctx) ccl = Term.it_mkProd_or_LetIn ccl ctx @@ -946,6 +954,7 @@ let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation = let len = List.length binders_rel in let top_env = push_rel_context binders_rel env in let top_arity = interp_type_evars top_env evdref arityc in + let top_arity = EConstr.Unsafe.to_constr top_arity in let full_arity = Term.it_mkProd_or_LetIn top_arity binders_rel in let argtyp, letbinders, make = telescope binders_rel in let argname = Id.of_string "recarg" in @@ -953,22 +962,24 @@ let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation = let binders = letbinders @ [arg] in let binders_env = push_rel_context binders_rel env in let rel, _ = interp_constr_evars_impls env evdref r in - let relty = Typing.unsafe_type_of env !evdref (EConstr.of_constr rel) in + let relty = Typing.unsafe_type_of env !evdref rel in let relargty = let error () = user_err ~loc:(constr_loc r) ~hdr:"Command.build_wellfounded" - (Printer.pr_constr_env env !evdref rel ++ str " is not an homogeneous binary relation.") + (Printer.pr_econstr_env env !evdref rel ++ str " is not an homogeneous binary relation.") in try - let ctx, ar = Reductionops.splay_prod_n env !evdref 2 (EConstr.of_constr relty) in + let ctx, ar = Reductionops.splay_prod_n env !evdref 2 relty in match ctx, EConstr.kind !evdref ar with | [LocalAssum (_,t); LocalAssum (_,u)], Sort (Prop Null) when Reductionops.is_conv env !evdref (EConstr.of_constr t) (EConstr.of_constr u) -> t | _, _ -> error () with e when CErrors.noncritical e -> error () in + let rel = EConstr.Unsafe.to_constr rel in let measure = interp_casted_constr_evars binders_env evdref measure relargty in + let measure = EConstr.Unsafe.to_constr measure in let wf_rel, wf_rel_fun, measure_fn = let measure_body, measure = it_mkLambda_or_LetIn measure letbinders, @@ -1025,6 +1036,7 @@ let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation = interp_casted_constr_evars (push_rel_context ctx env) evdref ~impls:newimpls body (lift 1 top_arity) in + let intern_body = EConstr.Unsafe.to_constr intern_body in let intern_body_lam = it_mkLambda_or_LetIn intern_body (curry_fun :: lift_lets @ fun_bl) in let prop = mkLambda (Name argname, argtyp, top_arity_let) in let def = @@ -1035,6 +1047,7 @@ let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation = prop |]) in let def = Typing.e_solve_evars env evdref (EConstr.of_constr def) in + let def = EConstr.Unsafe.to_constr def in let _ = evdref := Evarutil.nf_evar_map !evdref in let def = mkApp (def, [|intern_body_lam|]) in let binders_rel = nf_evar_context !evdref binders_rel in @@ -1110,7 +1123,7 @@ let interp_recursive isfix fixl notations = let fixprot = try let app = mkApp (delayed_force fix_proto, [|sort; t|]) in - Typing.e_solve_evars env evdref (EConstr.of_constr app) + EConstr.Unsafe.to_constr (Typing.e_solve_evars env evdref (EConstr.of_constr app)) with e when CErrors.noncritical e -> t in LocalAssum (id,fixprot) :: env' diff --git a/toplevel/record.ml b/toplevel/record.ml index 73035deb0..a65f5252e 100644 --- a/toplevel/record.ml +++ b/toplevel/record.ml @@ -66,7 +66,8 @@ let interp_fields_evars env evars impls_env nots l = List.fold_left2 (fun (env, uimpls, params, impls) no ((loc, i), b, t) -> let t', impl = interp_type_evars_impls env evars ~impls t in - let b' = Option.map (fun x -> fst (interp_casted_constr_evars_impls env evars ~impls x t')) b in + let b' = Option.map (fun x -> EConstr.Unsafe.to_constr (fst (interp_casted_constr_evars_impls env evars ~impls x t'))) b in + let t' = EConstr.Unsafe.to_constr t' in let impls = match i with | Anonymous -> impls @@ -120,7 +121,8 @@ let typecheck_params_and_fields def id pl t ps nots fs = match t with | CSort (_, Misctypes.GType []) -> true | _ -> false in let s = interp_type_evars env evars ~impls:empty_internalization_env t in - let sred = Reductionops.whd_all env !evars (EConstr.of_constr s) in + let sred = Reductionops.whd_all env !evars s in + let s = EConstr.Unsafe.to_constr s in let sred = EConstr.Unsafe.to_constr sred in (match kind_of_term sred with | Sort s' -> diff --git a/toplevel/vernacentries.ml b/toplevel/vernacentries.ml index ee2aacfc5..4376f51dc 100644 --- a/toplevel/vernacentries.ml +++ b/toplevel/vernacentries.ml @@ -108,7 +108,7 @@ let show_intro all = let {Evd.it=gls ; sigma=sigma; } = Proof.V82.subgoals pf in if not (List.is_empty gls) then begin let gl = {Evd.it=List.hd gls ; sigma = sigma; } in - let l,_= decompose_prod_assum sigma (Termops.strip_outer_cast sigma (EConstr.of_constr (pf_concl gl))) in + let l,_= decompose_prod_assum sigma (Termops.strip_outer_cast sigma (pf_concl gl)) in if all then let lid = Tactics.find_intro_names l gl in Feedback.msg_notice (hov 0 (prlist_with_sep spc pr_id lid)) @@ -1572,6 +1572,7 @@ let get_current_context_of_args = function let vernac_check_may_eval redexp glopt rc = let (sigma, env) = get_current_context_of_args glopt in let sigma', c = interp_open_constr env sigma rc in + let c = EConstr.Unsafe.to_constr c in let sigma' = Evarconv.consider_remaining_unif_problems env sigma' in Evarconv.check_problems_are_solved env sigma'; let sigma',nf = Evarutil.nf_evars_and_universes sigma' in -- cgit v1.2.3 From 0c56c953670d69f40e9554e35bdb206c2fb80911 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 25 Nov 2016 22:40:26 +0100 Subject: Micromega API using EConstr. --- plugins/micromega/coq_micromega.ml | 272 ++++++++++++++++++++----------------- 1 file changed, 144 insertions(+), 128 deletions(-) diff --git a/plugins/micromega/coq_micromega.ml b/plugins/micromega/coq_micromega.ml index a2346cc90..82218a35c 100644 --- a/plugins/micromega/coq_micromega.ml +++ b/plugins/micromega/coq_micromega.ml @@ -20,6 +20,8 @@ open Pp open Mutils open Goptions +module Term = EConstr + (** * Debug flag *) @@ -330,6 +332,8 @@ struct open Coqlib open Term + open Constr + open EConstr (** * Location of the Coq libraries. @@ -373,6 +377,7 @@ struct * ZMicromega.v *) + let gen_constant_in_modules s m n = EConstr.of_constr (gen_constant_in_modules s m n) let init_constant = gen_constant_in_modules "ZMicromega" init_modules let constant = gen_constant_in_modules "ZMicromega" coq_modules let bin_constant = gen_constant_in_modules "ZMicromega" bin_module @@ -599,12 +604,12 @@ struct (* A simple but useful getter function *) - let get_left_construct term = - match Term.kind_of_term term with - | Term.Construct((_,i),_) -> (i,[| |]) - | Term.App(l,rst) -> - (match Term.kind_of_term l with - | Term.Construct((_,i),_) -> (i,rst) + let get_left_construct sigma term = + match EConstr.kind sigma term with + | Constr.Construct((_,i),_) -> (i,[| |]) + | Constr.App(l,rst) -> + (match EConstr.kind sigma l with + | Constr.Construct((_,i),_) -> (i,rst) | _ -> raise ParseError ) | _ -> raise ParseError @@ -613,11 +618,11 @@ struct (* parse/dump/print from numbers up to expressions and formulas *) - let rec parse_nat term = - let (i,c) = get_left_construct term in + let rec parse_nat sigma term = + let (i,c) = get_left_construct sigma term in match i with | 1 -> Mc.O - | 2 -> Mc.S (parse_nat (c.(0))) + | 2 -> Mc.S (parse_nat sigma (c.(0))) | i -> raise ParseError let pp_nat o n = Printf.fprintf o "%i" (CoqToCaml.nat n) @@ -627,11 +632,11 @@ struct | Mc.O -> Lazy.force coq_O | Mc.S p -> Term.mkApp(Lazy.force coq_S,[| dump_nat p |]) - let rec parse_positive term = - let (i,c) = get_left_construct term in + let rec parse_positive sigma term = + let (i,c) = get_left_construct sigma term in match i with - | 1 -> Mc.XI (parse_positive c.(0)) - | 2 -> Mc.XO (parse_positive c.(0)) + | 1 -> Mc.XI (parse_positive sigma c.(0)) + | 2 -> Mc.XO (parse_positive sigma c.(0)) | 3 -> Mc.XH | i -> raise ParseError @@ -661,12 +666,12 @@ struct let dump_pair t1 t2 dump_t1 dump_t2 (x,y) = Term.mkApp(Lazy.force coq_pair,[| t1 ; t2 ; dump_t1 x ; dump_t2 y|]) - let parse_z term = - let (i,c) = get_left_construct term in + let parse_z sigma term = + let (i,c) = get_left_construct sigma term in match i with | 1 -> Mc.Z0 - | 2 -> Mc.Zpos (parse_positive c.(0)) - | 3 -> Mc.Zneg (parse_positive c.(0)) + | 2 -> Mc.Zpos (parse_positive sigma c.(0)) + | 3 -> Mc.Zneg (parse_positive sigma c.(0)) | i -> raise ParseError let dump_z x = @@ -686,10 +691,10 @@ struct Term.mkApp(Lazy.force coq_Qmake, [| dump_z q.Micromega.qnum ; dump_positive q.Micromega.qden|]) - let parse_q term = - match Term.kind_of_term term with - | Term.App(c, args) -> if Constr.equal c (Lazy.force coq_Qmake) then - {Mc.qnum = parse_z args.(0) ; Mc.qden = parse_positive args.(1) } + let parse_q sigma term = + match EConstr.kind sigma term with + | Constr.App(c, args) -> if EConstr.eq_constr sigma c (Lazy.force coq_Qmake) then + {Mc.qnum = parse_z sigma args.(0) ; Mc.qden = parse_positive sigma args.(1) } else raise ParseError | _ -> raise ParseError @@ -719,27 +724,27 @@ struct | Mc.CInv t -> Term.mkApp(Lazy.force coq_CInv, [| dump_Rcst t |]) | Mc.COpp t -> Term.mkApp(Lazy.force coq_COpp, [| dump_Rcst t |]) - let rec parse_Rcst term = - let (i,c) = get_left_construct term in + let rec parse_Rcst sigma term = + let (i,c) = get_left_construct sigma term in match i with | 1 -> Mc.C0 | 2 -> Mc.C1 - | 3 -> Mc.CQ (parse_q c.(0)) - | 4 -> Mc.CPlus(parse_Rcst c.(0), parse_Rcst c.(1)) - | 5 -> Mc.CMinus(parse_Rcst c.(0), parse_Rcst c.(1)) - | 6 -> Mc.CMult(parse_Rcst c.(0), parse_Rcst c.(1)) - | 7 -> Mc.CInv(parse_Rcst c.(0)) - | 8 -> Mc.COpp(parse_Rcst c.(0)) + | 3 -> Mc.CQ (parse_q sigma c.(0)) + | 4 -> Mc.CPlus(parse_Rcst sigma c.(0), parse_Rcst sigma c.(1)) + | 5 -> Mc.CMinus(parse_Rcst sigma c.(0), parse_Rcst sigma c.(1)) + | 6 -> Mc.CMult(parse_Rcst sigma c.(0), parse_Rcst sigma c.(1)) + | 7 -> Mc.CInv(parse_Rcst sigma c.(0)) + | 8 -> Mc.COpp(parse_Rcst sigma c.(0)) | _ -> raise ParseError - let rec parse_list parse_elt term = - let (i,c) = get_left_construct term in + let rec parse_list sigma parse_elt term = + let (i,c) = get_left_construct sigma term in match i with | 1 -> [] - | 2 -> parse_elt c.(1) :: parse_list parse_elt c.(2) + | 2 -> parse_elt sigma c.(1) :: parse_list sigma parse_elt c.(2) | i -> raise ParseError let rec dump_list typ dump_elt l = @@ -872,9 +877,9 @@ struct dump_op o ; dump_expr typ dump_constant e2|]) - let assoc_const x l = + let assoc_const sigma x l = try - snd (List.find (fun (x',y) -> Constr.equal x (Lazy.force x')) l) + snd (List.find (fun (x',y) -> EConstr.eq_constr sigma x (Lazy.force x')) l) with Not_found -> raise ParseError @@ -898,35 +903,37 @@ struct let has_typ gl t1 typ = let ty = Retyping.get_type_of (Tacmach.pf_env gl) (Tacmach.project gl) t1 in - EConstr.eq_constr (Tacmach.project gl) ty (EConstr.of_constr typ) + EConstr.eq_constr (Tacmach.project gl) ty typ let is_convertible gl t1 t2 = - Reductionops.is_conv (Tacmach.pf_env gl) (Tacmach.project gl) (EConstr.of_constr t1) (EConstr.of_constr t2) + Reductionops.is_conv (Tacmach.pf_env gl) (Tacmach.project gl) t1 t2 let parse_zop gl (op,args) = - match kind_of_term op with - | Const (x,_) -> (assoc_const op zop_table, args.(0) , args.(1)) + let sigma = Tacmach.project gl in + match EConstr.kind sigma op with + | Const (x,_) -> (assoc_const sigma op zop_table, args.(0) , args.(1)) | Ind((n,0),_) -> - if Constr.equal op (Lazy.force coq_Eq) && is_convertible gl args.(0) (Lazy.force coq_Z) + if EConstr.eq_constr sigma op (Lazy.force coq_Eq) && is_convertible gl args.(0) (Lazy.force coq_Z) then (Mc.OpEq, args.(1), args.(2)) else raise ParseError | _ -> failwith "parse_zop" let parse_rop gl (op,args) = - match kind_of_term op with - | Const (x,_) -> (assoc_const op rop_table, args.(0) , args.(1)) + let sigma = Tacmach.project gl in + match EConstr.kind sigma op with + | Const (x,_) -> (assoc_const sigma op rop_table, args.(0) , args.(1)) | Ind((n,0),_) -> - if Constr.equal op (Lazy.force coq_Eq) && is_convertible gl args.(0) (Lazy.force coq_R) + if EConstr.eq_constr sigma op (Lazy.force coq_Eq) && is_convertible gl args.(0) (Lazy.force coq_R) then (Mc.OpEq, args.(1), args.(2)) else raise ParseError | _ -> failwith "parse_zop" let parse_qop gl (op,args) = - (assoc_const op qop_table, args.(0) , args.(1)) + (assoc_const (Tacmach.project gl) op qop_table, args.(0) , args.(1)) - let is_constant t = (* This is an approx *) - match kind_of_term t with + let is_constant sigma t = (* This is an approx *) + match EConstr.kind sigma t with | Construct(i,_) -> true | _ -> false @@ -936,9 +943,9 @@ struct | Power | Ukn of string - let assoc_ops x l = + let assoc_ops sigma x l = try - snd (List.find (fun (x',y) -> Constr.equal x (Lazy.force x')) l) + snd (List.find (fun (x',y) -> EConstr.eq_constr sigma x (Lazy.force x')) l) with Not_found -> Ukn "Oups" @@ -950,12 +957,12 @@ struct struct type t = constr list - let compute_rank_add env v = + let compute_rank_add env sigma v = let rec _add env n v = match env with | [] -> ([v],n) | e::l -> - if eq_constr e v + if eq_constr sigma e v then (env,n) else let (env,n) = _add l ( n+1) v in @@ -963,13 +970,13 @@ struct let (env, n) = _add env 1 v in (env, CamlToCoq.positive n) - let get_rank env v = + let get_rank env sigma v = let rec _get_rank env n = match env with | [] -> raise (Invalid_argument "get_rank") | e::l -> - if eq_constr e v + if eq_constr sigma e v then n else _get_rank l (n+1) in _get_rank env 1 @@ -985,9 +992,9 @@ struct * This is the big generic function for expression parsers. *) - let parse_expr parse_constant parse_exp ops_spec env term = + let parse_expr sigma parse_constant parse_exp ops_spec env term = if debug - then Feedback.msg_debug (Pp.str "parse_expr: " ++ Printer.prterm term); + then Feedback.msg_debug (Pp.str "parse_expr: " ++ Printer.pr_leconstr term); (* let constant_or_variable env term = @@ -998,7 +1005,7 @@ struct (Mc.PEX n , env) in *) let parse_variable env term = - let (env,n) = Env.compute_rank_add env term in + let (env,n) = Env.compute_rank_add env sigma term in (Mc.PEX n , env) in let rec parse_expr env term = @@ -1009,12 +1016,12 @@ struct try (Mc.PEc (parse_constant term) , env) with ParseError -> - match kind_of_term term with + match EConstr.kind sigma term with | App(t,args) -> ( - match kind_of_term t with + match EConstr.kind sigma t with | Const c -> - ( match assoc_ops t ops_spec with + ( match assoc_ops sigma t ops_spec with | Binop f -> combine env f (args.(0),args.(1)) | Opp -> let (expr,env) = parse_expr env args.(0) in (Mc.PEopp expr, env) @@ -1026,12 +1033,12 @@ struct (power , env) with e when CErrors.noncritical e -> (* if the exponent is a variable *) - let (env,n) = Env.compute_rank_add env term in (Mc.PEX n, env) + let (env,n) = Env.compute_rank_add env sigma term in (Mc.PEX n, env) end | Ukn s -> if debug then (Printf.printf "unknown op: %s\n" s; flush stdout;); - let (env,n) = Env.compute_rank_add env term in (Mc.PEX n, env) + let (env,n) = Env.compute_rank_add env sigma term in (Mc.PEX n, env) ) | _ -> parse_variable env term ) @@ -1074,60 +1081,60 @@ struct (* coq_Rdiv , (fun x y -> Mc.CMult(x,Mc.CInv y)) ;*) ] - let rec rconstant term = - match Term.kind_of_term term with + let rec rconstant sigma term = + match EConstr.kind sigma term with | Const x -> - if Constr.equal term (Lazy.force coq_R0) + if EConstr.eq_constr sigma term (Lazy.force coq_R0) then Mc.C0 - else if Constr.equal term (Lazy.force coq_R1) + else if EConstr.eq_constr sigma term (Lazy.force coq_R1) then Mc.C1 else raise ParseError | App(op,args) -> begin try (* the evaluation order is important in the following *) - let f = assoc_const op rconst_assoc in - let a = rconstant args.(0) in - let b = rconstant args.(1) in + let f = assoc_const sigma op rconst_assoc in + let a = rconstant sigma args.(0) in + let b = rconstant sigma args.(1) in f a b with ParseError -> match op with - | op when Constr.equal op (Lazy.force coq_Rinv) -> - let arg = rconstant args.(0) in + | op when EConstr.eq_constr sigma op (Lazy.force coq_Rinv) -> + let arg = rconstant sigma args.(0) in if Mc.qeq_bool (Mc.q_of_Rcst arg) {Mc.qnum = Mc.Z0 ; Mc.qden = Mc.XH} then raise ParseError (* This is a division by zero -- no semantics *) else Mc.CInv(arg) - | op when Constr.equal op (Lazy.force coq_IQR) -> Mc.CQ (parse_q args.(0)) - | op when Constr.equal op (Lazy.force coq_IZR) -> Mc.CZ (parse_z args.(0)) + | op when EConstr.eq_constr sigma op (Lazy.force coq_IQR) -> Mc.CQ (parse_q sigma args.(0)) + | op when EConstr.eq_constr sigma op (Lazy.force coq_IZR) -> Mc.CZ (parse_z sigma args.(0)) | _ -> raise ParseError end | _ -> raise ParseError - let rconstant term = + let rconstant sigma term = if debug - then Feedback.msg_debug (Pp.str "rconstant: " ++ Printer.prterm term ++ fnl ()); - let res = rconstant term in + then Feedback.msg_debug (Pp.str "rconstant: " ++ Printer.pr_leconstr term ++ fnl ()); + let res = rconstant sigma term in if debug then (Printf.printf "rconstant -> %a\n" pp_Rcst res ; flush stdout) ; res - let parse_zexpr = parse_expr - zconstant + let parse_zexpr sigma = parse_expr sigma + (zconstant sigma) (fun expr x -> - let exp = (parse_z x) in + let exp = (parse_z sigma x) in match exp with | Mc.Zneg _ -> Mc.PEc Mc.Z0 | _ -> Mc.PEpow(expr, Mc.Z.to_N exp)) zop_spec - let parse_qexpr = parse_expr - qconstant + let parse_qexpr sigma = parse_expr sigma + (qconstant sigma) (fun expr x -> - let exp = parse_z x in + let exp = parse_z sigma x in match exp with | Mc.Zneg _ -> begin @@ -1139,21 +1146,22 @@ struct Mc.PEpow(expr,exp)) qop_spec - let parse_rexpr = parse_expr - rconstant + let parse_rexpr sigma = parse_expr sigma + (rconstant sigma) (fun expr x -> - let exp = Mc.N.of_nat (parse_nat x) in + let exp = Mc.N.of_nat (parse_nat sigma x) in Mc.PEpow(expr,exp)) rop_spec let parse_arith parse_op parse_expr env cstr gl = + let sigma = Tacmach.project gl in if debug - then Feedback.msg_debug (Pp.str "parse_arith: " ++ Printer.prterm cstr ++ fnl ()); - match kind_of_term cstr with + then Feedback.msg_debug (Pp.str "parse_arith: " ++ Printer.pr_leconstr cstr ++ fnl ()); + match EConstr.kind sigma cstr with | App(op,args) -> let (op,lhs,rhs) = parse_op gl (op,args) in - let (e1,env) = parse_expr env lhs in - let (e2,env) = parse_expr env rhs in + let (e1,env) = parse_expr sigma env lhs in + let (e2,env) = parse_expr sigma env rhs in ({Mc.flhs = e1; Mc.fop = op;Mc.frhs = e2},env) | _ -> failwith "error : parse_arith(2)" @@ -1191,6 +1199,7 @@ struct *) let parse_formula gl parse_atom env tg term = + let sigma = Tacmach.project gl in let parse_atom env tg t = try @@ -1199,34 +1208,34 @@ struct with e when CErrors.noncritical e -> (X(t),env,tg) in let is_prop term = - let sort = Retyping.get_sort_of (Tacmach.pf_env gl) (Tacmach.project gl) (EConstr.of_constr term) in - Term.is_prop_sort sort in + let sort = Retyping.get_sort_of (Tacmach.pf_env gl) (Tacmach.project gl) term in + Sorts.is_prop sort in let rec xparse_formula env tg term = - match kind_of_term term with + match EConstr.kind sigma term with | App(l,rst) -> (match rst with - | [|a;b|] when eq_constr l (Lazy.force coq_and) -> + | [|a;b|] when eq_constr sigma l (Lazy.force coq_and) -> let f,env,tg = xparse_formula env tg a in let g,env, tg = xparse_formula env tg b in mkformula_binary mkC term f g,env,tg - | [|a;b|] when eq_constr l (Lazy.force coq_or) -> + | [|a;b|] when eq_constr sigma l (Lazy.force coq_or) -> let f,env,tg = xparse_formula env tg a in let g,env,tg = xparse_formula env tg b in mkformula_binary mkD term f g,env,tg - | [|a|] when eq_constr l (Lazy.force coq_not) -> + | [|a|] when eq_constr sigma l (Lazy.force coq_not) -> let (f,env,tg) = xparse_formula env tg a in (N(f), env,tg) - | [|a;b|] when eq_constr l (Lazy.force coq_iff) -> + | [|a;b|] when eq_constr sigma l (Lazy.force coq_iff) -> let f,env,tg = xparse_formula env tg a in let g,env,tg = xparse_formula env tg b in mkformula_binary mkIff term f g,env,tg | _ -> parse_atom env tg term) - | Prod(typ,a,b) when EConstr.Vars.noccurn Evd.empty 1 (EConstr.of_constr b) (** FIXME *) -> + | Prod(typ,a,b) when Vars.noccurn sigma 1 b -> let f,env,tg = xparse_formula env tg a in let g,env,tg = xparse_formula env tg b in mkformula_binary mkI term f g,env,tg - | _ when eq_constr term (Lazy.force coq_True) -> (TT,env,tg) - | _ when eq_constr term (Lazy.force coq_False) -> (FF,env,tg) + | _ when eq_constr sigma term (Lazy.force coq_True) -> (TT,env,tg) + | _ when eq_constr sigma term (Lazy.force coq_False) -> (FF,env,tg) | _ when is_prop term -> X(term),env,tg | _ -> raise ParseError in @@ -1246,10 +1255,10 @@ struct xdump f - let prop_env_of_formula form = + let prop_env_of_formula sigma form = let rec doit env = function | TT | FF | A(_,_,_) -> env - | X t -> fst (Env.compute_rank_add env t) + | X t -> fst (Env.compute_rank_add env sigma t) | C(f1,f2) | D(f1,f2) | I(f1,_,f2) -> doit (doit env f1) f2 | N f -> doit env f in @@ -1380,14 +1389,22 @@ let dump_rexpr = lazy *) -let rec make_goal_of_formula dexpr form = +let prodn n env b = + let rec prodrec = function + | (0, env, b) -> b + | (n, ((v,t)::l), b) -> prodrec (n-1, l, mkProd (v,t,b)) + | _ -> assert false + in + prodrec (n,env,b) + +let make_goal_of_formula sigma dexpr form = let vars_idx = List.mapi (fun i v -> (v, i+1)) (ISet.elements (var_env_of_formula form)) in (* List.iter (fun (v,i) -> Printf.fprintf stdout "var %i has index %i\n" v i) vars_idx ;*) - let props = prop_env_of_formula form in + let props = prop_env_of_formula sigma form in let vars_n = List.map (fun (_,i) -> (Names.id_of_string (Printf.sprintf "__x%i" i)) , dexpr.interp_typ) vars_idx in let props_n = List.mapi (fun i _ -> (Names.id_of_string (Printf.sprintf "__p%i" (i+1))) , Term.mkProp) props in @@ -1428,7 +1445,7 @@ let rec make_goal_of_formula dexpr form = | I(x,_,y) -> mkArrow (xdump pi xi x) (xdump (pi+1) (xi+1) y) | N(x) -> mkArrow (xdump pi xi x) (Lazy.force coq_False) | A(x,_,_) -> dump_cstr xi x - | X(t) -> let idx = Env.get_rank props t in + | X(t) -> let idx = Env.get_rank props sigma t in mkRel (pi+idx) in let nb_vars = List.length vars_n in @@ -1437,13 +1454,13 @@ let rec make_goal_of_formula dexpr form = (* Printf.fprintf stdout "NBProps : %i\n" nb_props ;*) let subst_prop p = - let idx = Env.get_rank props p in + let idx = Env.get_rank props sigma p in mkVar (Names.id_of_string (Printf.sprintf "__p%i" idx)) in let form' = map_prop subst_prop form in - (Term.prodn nb_props (List.map (fun (x,y) -> Names.Name x,y) props_n) - (Term.prodn nb_vars (List.map (fun (x,y) -> Names.Name x,y) vars_n) + (prodn nb_props (List.map (fun (x,y) -> Names.Name x,y) props_n) + (prodn nb_vars (List.map (fun (x,y) -> Names.Name x,y) vars_n) (xdump (List.length vars_n) 0 form)), List.rev props_n, List.rev var_name_pos,form') @@ -1461,7 +1478,7 @@ let rec make_goal_of_formula dexpr form = xset (Term.mkNamedLetIn (Names.Id.of_string name) expr typ acc) l in - EConstr.of_constr (xset concl l) + xset concl l end (** * MODULE END: M @@ -1518,17 +1535,17 @@ let rec apply_ids t ids = | i::ids -> apply_ids (Term.mkApp(t,[| Term.mkVar i |])) ids let coq_Node = - (Coqlib.gen_constant_in_modules "VarMap" + EConstr.of_constr (Coqlib.gen_constant_in_modules "VarMap" [["Coq" ; "micromega" ; "VarMap"];["VarMap"]] "Node") let coq_Leaf = - (Coqlib.gen_constant_in_modules "VarMap" + EConstr.of_constr (Coqlib.gen_constant_in_modules "VarMap" [["Coq" ; "micromega" ; "VarMap"];["VarMap"]] "Leaf") let coq_Empty = - (Coqlib.gen_constant_in_modules "VarMap" + EConstr.of_constr (Coqlib.gen_constant_in_modules "VarMap" [["Coq" ; "micromega" ;"VarMap"];["VarMap"]] "Empty") let coq_VarMap = - (Coqlib.gen_constant_in_modules "VarMap" + EConstr.of_constr (Coqlib.gen_constant_in_modules "VarMap" [["Coq" ; "micromega" ; "VarMap"] ; ["VarMap"]] "t") @@ -1606,7 +1623,6 @@ let rec parse_hyps gl parse_arith env tg hyps = match hyps with | [] -> ([],env,tg) | (i,t)::l -> - let t = EConstr.Unsafe.to_constr t in let (lhyps,env,tg) = parse_hyps gl parse_arith env tg l in try let (c,env,tg) = parse_formula gl parse_arith env tg t in @@ -1714,7 +1730,7 @@ let micromega_order_change spec cert cert_typ env ff (*: unit Proofview.tactic* ("__varmap", vm, Term.mkApp( coq_VarMap, [|spec.typ|])); ("__wit", cert, cert_typ) ] - (EConstr.Unsafe.to_constr (Tacmach.pf_concl gl))) + (Tacmach.pf_concl gl)) ] end } @@ -1905,7 +1921,7 @@ let micromega_tauto negate normalise unsat deduce spec prover env polys1 polys2 let formula_typ = (Term.mkApp(Lazy.force coq_Cstr, [|spec.coeff|])) in let ff = dump_formula formula_typ (dump_cstr spec.typ spec.dump_coeff) ff in - Feedback.msg_notice (Printer.prterm ff); + Feedback.msg_notice (Printer.pr_leconstr ff); Printf.fprintf stdout "cnf : %a\n" (pp_cnf (fun o _ -> ())) cnf_ff end; @@ -1930,7 +1946,7 @@ let micromega_tauto negate normalise unsat deduce spec prover env polys1 polys2 let formula_typ = (Term.mkApp( Lazy.force coq_Cstr,[| spec.coeff|])) in let ff' = dump_formula formula_typ (dump_cstr spec.typ spec.dump_coeff) ff' in - Feedback.msg_notice (Printer.prterm ff'); + Feedback.msg_notice (Printer.pr_leconstr ff'); Printf.fprintf stdout "cnf : %a\n" (pp_cnf (fun o _ -> ())) cnf_ff' end; @@ -1964,8 +1980,8 @@ let micromega_gen spec dumpexpr prover tac = Proofview.Goal.nf_enter { enter = begin fun gl -> let gl = Tacmach.New.of_old (fun x -> x) gl in + let sigma = Tacmach.project gl in let concl = Tacmach.pf_concl gl in - let concl = EConstr.Unsafe.to_constr concl in let hyps = Tacmach.pf_hyps_types gl in try let (hyps,concl,env) = parse_goal gl parse_arith Env.empty hyps concl in @@ -1976,7 +1992,7 @@ let micromega_gen match micromega_tauto negate normalise unsat deduce spec prover env hyps concl gl with | None -> Tacticals.New.tclFAIL 0 (Pp.str " Cannot find witness") | Some (ids,ff',res') -> - let (arith_goal,props,vars,ff_arith) = make_goal_of_formula dumpexpr ff' in + let (arith_goal,props,vars,ff_arith) = make_goal_of_formula sigma dumpexpr ff' in let intro (id,_) = Tactics.introduction id in let intro_vars = Tacticals.New.tclTHENLIST (List.map intro vars) in @@ -1989,7 +2005,7 @@ let micromega_gen micromega_order_change spec res' (Term.mkApp(Lazy.force coq_list, [|spec.proof_typ|])) env' ff_arith ] in - let goal_props = List.rev (prop_env_of_formula ff') in + let goal_props = List.rev (prop_env_of_formula sigma ff') in let goal_vars = List.map (fun (_,i) -> List.nth env (i-1)) vars in @@ -2002,12 +2018,12 @@ let micromega_gen (Tacticals.New.tclTHEN tac_arith tac)) in Tacticals.New.tclTHENS - (Tactics.forward true (Some None) (ipat_of_name goal_name) (EConstr.of_constr arith_goal)) + (Tactics.forward true (Some None) (ipat_of_name goal_name) arith_goal) [ kill_arith; (Tacticals.New.tclTHENLIST - [(Tactics.generalize (List.map EConstr.mkVar ids)); - Tactics.exact_check (EConstr.of_constr (Term.applist (Term.mkVar goal_name, arith_args))) + [(Tactics.generalize (List.map Term.mkVar ids)); + Tactics.exact_check (Term.applist (Term.mkVar goal_name, arith_args)) ] ) ] with @@ -2049,11 +2065,11 @@ let micromega_order_changer cert env ff = [ ("__ff", ff, Term.mkApp(Lazy.force coq_Formula, [|formula_typ |])); ("__varmap", vm, Term.mkApp - (Coqlib.gen_constant_in_modules "VarMap" - [["Coq" ; "micromega" ; "VarMap"] ; ["VarMap"]] "t", [|typ|])); + (EConstr.of_constr (Coqlib.gen_constant_in_modules "VarMap" + [["Coq" ; "micromega" ; "VarMap"] ; ["VarMap"]] "t"), [|typ|])); ("__wit", cert, cert_typ) ] - (EConstr.Unsafe.to_constr (Tacmach.pf_concl gl)))); + (Tacmach.pf_concl gl))); (* Tacticals.New.tclTHENLIST (List.map (fun id -> (Tactics.introduction id)) ids)*) ] end } @@ -2073,8 +2089,8 @@ let micromega_genr prover tac = } in Proofview.Goal.nf_enter { enter = begin fun gl -> let gl = Tacmach.New.of_old (fun x -> x) gl in + let sigma = Tacmach.project gl in let concl = Tacmach.pf_concl gl in - let concl = EConstr.Unsafe.to_constr concl in let hyps = Tacmach.pf_hyps_types gl in try @@ -2092,7 +2108,7 @@ let micromega_genr prover tac = (List.filter (fun (n,_) -> List.mem n ids) hyps) concl in let ff' = abstract_wrt_formula ff' ff in - let (arith_goal,props,vars,ff_arith) = make_goal_of_formula (Lazy.force dump_rexpr) ff' in + let (arith_goal,props,vars,ff_arith) = make_goal_of_formula sigma (Lazy.force dump_rexpr) ff' in let intro (id,_) = Tactics.introduction id in let intro_vars = Tacticals.New.tclTHENLIST (List.map intro vars) in @@ -2104,7 +2120,7 @@ let micromega_genr prover tac = let tac_arith = Tacticals.New.tclTHENLIST [ intro_props ; intro_vars ; micromega_order_changer res' env' ff_arith ] in - let goal_props = List.rev (prop_env_of_formula ff') in + let goal_props = List.rev (prop_env_of_formula sigma ff') in let goal_vars = List.map (fun (_,i) -> List.nth env (i-1)) vars in @@ -2117,12 +2133,12 @@ let micromega_genr prover tac = (Tacticals.New.tclTHEN tac_arith tac)) in Tacticals.New.tclTHENS - (Tactics.forward true (Some None) (ipat_of_name goal_name) (EConstr.of_constr arith_goal)) + (Tactics.forward true (Some None) (ipat_of_name goal_name) arith_goal) [ kill_arith; (Tacticals.New.tclTHENLIST - [(Tactics.generalize (List.map EConstr.mkVar ids)); - Tactics.exact_check (EConstr.of_constr (Term.applist (Term.mkVar goal_name, arith_args))) + [(Tactics.generalize (List.map Term.mkVar ids)); + Tactics.exact_check (Term.applist (Term.mkVar goal_name, arith_args)) ] ) ] -- cgit v1.2.3 From e1010899051546467b790bca0409174bde824270 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sat, 26 Nov 2016 00:36:34 +0100 Subject: Omega API using EConstr. --- plugins/omega/coq_omega.ml | 252 ++++++++++++++++++++++++--------------------- 1 file changed, 136 insertions(+), 116 deletions(-) diff --git a/plugins/omega/coq_omega.ml b/plugins/omega/coq_omega.ml index 665486272..9e0d591b6 100644 --- a/plugins/omega/coq_omega.ml +++ b/plugins/omega/coq_omega.ml @@ -18,6 +18,7 @@ open Util open Names open Nameops open Term +open EConstr open Tacticals open Tacmach open Tactics @@ -149,7 +150,7 @@ let mk_then = tclTHENLIST let exists_tac c = constructor_tac false (Some 1) 1 (ImplicitBindings [c]) -let generalize_tac t = generalize (List.map EConstr.of_constr t) +let generalize_tac t = generalize t let elim t = simplest_elim t let exact t = Tacmach.refine t let unfold s = Tactics.unfold_in_concl [Locus.AllOccurrences, Lazy.force s] @@ -172,8 +173,8 @@ let tag_hypothesis,tag_of_hyp, hyp_of_tag, clear_tags = let hide_constr,find_constr,clear_constr_tables,dump_tables = let l = ref ([]:(constr * (Id.t * Id.t * bool)) list) in (fun h id eg b -> l := (h,(id,eg,b)):: !l), - (fun h -> - try List.assoc_f eq_constr_nounivs h !l with Not_found -> failwith "find_contr"), + (fun sigma h -> + try List.assoc_f (eq_constr_nounivs sigma) h !l with Not_found -> failwith "find_contr"), (fun () -> l := []), (fun () -> !l) @@ -197,6 +198,7 @@ let coq_modules = init_modules @arith_modules @ [logic_dir] @ zarith_base_modules @ [["Coq"; "omega"; "OmegaLemmas"]] +let gen_constant_in_modules n m s = EConstr.of_constr (gen_constant_in_modules n m s) let init_constant = gen_constant_in_modules "Omega" init_modules let constant = gen_constant_in_modules "Omega" coq_modules @@ -348,11 +350,18 @@ let coq_not_iff = lazy (constant "not_iff") let coq_not_not = lazy (constant "not_not") let coq_imp_simp = lazy (constant "imp_simp") let coq_iff = lazy (constant "iff") +let coq_not = lazy (init_constant "not") +let coq_and = lazy (init_constant "and") +let coq_or = lazy (init_constant "or") +let coq_eq = lazy (init_constant "eq") +let coq_ex = lazy (init_constant "ex") +let coq_False = lazy (init_constant "False") +let coq_True = lazy (init_constant "True") (* uses build_coq_and, build_coq_not, build_coq_or, build_coq_ex *) (* For unfold *) -let evaluable_ref_of_constr s c = match kind_of_term (Lazy.force c) with +let evaluable_ref_of_constr s c = match EConstr.kind Evd.empty (Lazy.force c) with | Const (kn,u) when Tacred.is_evaluable (Global.env()) (EvalConstRef kn) -> EvalConstRef kn | _ -> anomaly ~label:"Coq_omega" (Pp.str (s^" is not an evaluable constant")) @@ -364,21 +373,21 @@ let sp_Zle = lazy (evaluable_ref_of_constr "Z.le" coq_Zle) let sp_Zgt = lazy (evaluable_ref_of_constr "Z.gt" coq_Zgt) let sp_Zge = lazy (evaluable_ref_of_constr "Z.ge" coq_Zge) let sp_Zlt = lazy (evaluable_ref_of_constr "Z.lt" coq_Zlt) -let sp_not = lazy (evaluable_ref_of_constr "not" (lazy (build_coq_not ()))) +let sp_not = lazy (evaluable_ref_of_constr "not" coq_not) let mk_var v = mkVar (Id.of_string v) let mk_plus t1 t2 = mkApp (Lazy.force coq_Zplus, [| t1; t2 |]) let mk_times t1 t2 = mkApp (Lazy.force coq_Zmult, [| t1; t2 |]) let mk_minus t1 t2 = mkApp (Lazy.force coq_Zminus, [| t1;t2 |]) -let mk_eq t1 t2 = mkApp (Universes.constr_of_global (build_coq_eq ()), +let mk_eq t1 t2 = mkApp (Lazy.force coq_eq, [| Lazy.force coq_Z; t1; t2 |]) let mk_le t1 t2 = mkApp (Lazy.force coq_Zle, [| t1; t2 |]) -let mk_gt t1 t2 = EConstr.of_constr (mkApp (Lazy.force coq_Zgt, [| t1; t2 |])) +let mk_gt t1 t2 = mkApp (Lazy.force coq_Zgt, [| t1; t2 |]) let mk_inv t = mkApp (Lazy.force coq_Zopp, [| t |]) -let mk_and t1 t2 = mkApp (build_coq_and (), [| t1; t2 |]) -let mk_or t1 t2 = mkApp (build_coq_or (), [| t1; t2 |]) -let mk_not t = mkApp (build_coq_not (), [| t |]) -let mk_eq_rel t1 t2 = mkApp (Universes.constr_of_global (build_coq_eq ()), +let mk_and t1 t2 = mkApp (Lazy.force coq_and, [| t1; t2 |]) +let mk_or t1 t2 = mkApp (Lazy.force coq_or, [| t1; t2 |]) +let mk_not t = mkApp (Lazy.force coq_not, [| t |]) +let mk_eq_rel t1 t2 = mkApp (Lazy.force coq_eq, [| Lazy.force coq_comparison; t1; t2 |]) let mk_inj t = mkApp (Lazy.force coq_Z_of_nat, [| t |]) @@ -420,22 +429,23 @@ type result = the term parts that we manipulate, but rather Var's. Said otherwise: all constr manipulated here are closed *) -let destructurate_prop t = - let c, args = decompose_app t in - match kind_of_term c, args with - | _, [_;_;_] when is_global (build_coq_eq ()) c -> Kapp (Eq,args) +let destructurate_prop sigma t = + let eq_constr c1 c2 = eq_constr sigma c1 c2 in + let c, args = decompose_app sigma t in + match EConstr.kind sigma c, args with + | _, [_;_;_] when eq_constr (Lazy.force coq_eq) c -> Kapp (Eq,args) | _, [_;_] when eq_constr c (Lazy.force coq_neq) -> Kapp (Neq,args) | _, [_;_] when eq_constr c (Lazy.force coq_Zne) -> Kapp (Zne,args) | _, [_;_] when eq_constr c (Lazy.force coq_Zle) -> Kapp (Zle,args) | _, [_;_] when eq_constr c (Lazy.force coq_Zlt) -> Kapp (Zlt,args) | _, [_;_] when eq_constr c (Lazy.force coq_Zge) -> Kapp (Zge,args) | _, [_;_] when eq_constr c (Lazy.force coq_Zgt) -> Kapp (Zgt,args) - | _, [_;_] when eq_constr c (build_coq_and ()) -> Kapp (And,args) - | _, [_;_] when eq_constr c (build_coq_or ()) -> Kapp (Or,args) + | _, [_;_] when eq_constr c (Lazy.force coq_and) -> Kapp (And,args) + | _, [_;_] when eq_constr c (Lazy.force coq_or) -> Kapp (Or,args) | _, [_;_] when eq_constr c (Lazy.force coq_iff) -> Kapp (Iff, args) - | _, [_] when eq_constr c (build_coq_not ()) -> Kapp (Not,args) - | _, [] when eq_constr c (build_coq_False ()) -> Kapp (False,args) - | _, [] when eq_constr c (build_coq_True ()) -> Kapp (True,args) + | _, [_] when eq_constr c (Lazy.force coq_not) -> Kapp (Not,args) + | _, [] when eq_constr c (Lazy.force coq_False) -> Kapp (False,args) + | _, [] when eq_constr c (Lazy.force coq_True) -> Kapp (True,args) | _, [_;_] when eq_constr c (Lazy.force coq_le) -> Kapp (Le,args) | _, [_;_] when eq_constr c (Lazy.force coq_lt) -> Kapp (Lt,args) | _, [_;_] when eq_constr c (Lazy.force coq_ge) -> Kapp (Ge,args) @@ -451,16 +461,18 @@ let destructurate_prop t = | Prod (Name _,_,_),[] -> error "Omega: Not a quantifier-free goal" | _ -> Kufo -let destructurate_type t = - let c, args = decompose_app t in - match kind_of_term c, args with +let destructurate_type sigma t = + let eq_constr c1 c2 = eq_constr sigma c1 c2 in + let c, args = decompose_app sigma t in + match EConstr.kind sigma c, args with | _, [] when eq_constr c (Lazy.force coq_Z) -> Kapp (Z,args) | _, [] when eq_constr c (Lazy.force coq_nat) -> Kapp (Nat,args) | _ -> Kufo -let destructurate_term t = - let c, args = decompose_app t in - match kind_of_term c, args with +let destructurate_term sigma t = + let eq_constr c1 c2 = eq_constr sigma c1 c2 in + let c, args = decompose_app sigma t in + match EConstr.kind sigma c, args with | _, [_;_] when eq_constr c (Lazy.force coq_Zplus) -> Kapp (Zplus,args) | _, [_;_] when eq_constr c (Lazy.force coq_Zmult) -> Kapp (Zmult,args) | _, [_;_] when eq_constr c (Lazy.force coq_Zminus) -> Kapp (Zminus,args) @@ -480,15 +492,16 @@ let destructurate_term t = | Var id,[] -> Kvar id | _ -> Kufo -let recognize_number t = +let recognize_number sigma t = + let eq_constr c1 c2 = eq_constr sigma c1 c2 in let rec loop t = - match decompose_app t with + match decompose_app sigma t with | f, [t] when eq_constr f (Lazy.force coq_xI) -> one + two * loop t | f, [t] when eq_constr f (Lazy.force coq_xO) -> two * loop t | f, [] when eq_constr f (Lazy.force coq_xH) -> one | _ -> failwith "not a number" in - match decompose_app t with + match decompose_app sigma t with | f, [t] when eq_constr f (Lazy.force coq_Zpos) -> loop t | f, [t] when eq_constr f (Lazy.force coq_Zneg) -> neg (loop t) | f, [] when eq_constr f (Lazy.force coq_Z0) -> zero @@ -504,9 +517,9 @@ type constr_path = | P_ARITY | P_ARG -let context operation path (t : constr) = +let context sigma operation path (t : constr) = let rec loop i p0 t = - match (p0,kind_of_term t) with + match (p0,EConstr.kind sigma t) with | (p, Cast (c,k,t)) -> mkCast (loop i p c,k,t) | ([], _) -> operation i t | ((P_APP n :: p), App (f,v)) -> @@ -517,7 +530,7 @@ let context operation path (t : constr) = let v' = Array.copy v in v'.(n) <- loop i p v'.(n); (mkCase (ci,q,c,v')) | ((P_ARITY :: p), App (f,l)) -> - appvect (loop i p f,l) + mkApp (loop i p f,l) | ((P_ARG :: p), App (f,v)) -> let v' = Array.copy v in v'.(0) <- loop i p v'.(0); mkApp (f,v') @@ -542,8 +555,8 @@ let context operation path (t : constr) = in loop 1 path t -let occurrence path (t : constr) = - let rec loop p0 t = match (p0,kind_of_term t) with +let occurrence sigma path (t : constr) = + let rec loop p0 t = match (p0,EConstr.kind sigma t) with | (p, Cast (c,_,_)) -> loop p c | ([], _) -> t | ((P_APP n :: p), App (f,v)) -> loop p v.(pred n) @@ -562,14 +575,13 @@ let occurrence path (t : constr) = in loop path t -let abstract_path typ path t = +let abstract_path sigma typ path t = let term_occur = ref (mkRel 0) in - let abstract = context (fun i t -> term_occur:= t; mkRel i) path t in + let abstract = context sigma (fun i t -> term_occur:= t; mkRel i) path t in mkLambda (Name (Id.of_string "x"), typ, abstract), !term_occur let focused_simpl path gl = - let newc = context (fun i t -> EConstr.Unsafe.to_constr (pf_nf gl (EConstr.of_constr t))) (List.rev path) (EConstr.Unsafe.to_constr (pf_concl gl)) in - let newc = EConstr.of_constr newc in + let newc = context (project gl) (fun i t -> pf_nf gl t) (List.rev path) (pf_concl gl) in Proofview.V82.of_tactic (convert_concl_no_check newc DEFAULTcast) gl let focused_simpl path = focused_simpl path @@ -632,8 +644,7 @@ let mkNewMeta () = mkMeta (Evarutil.new_meta()) let clever_rewrite_base_poly typ p result theorem gl = let full = pf_concl gl in - let full = EConstr.Unsafe.to_constr full in - let (abstracted,occ) = abstract_path typ (List.rev p) full in + let (abstracted,occ) = abstract_path (project gl) typ (List.rev p) full in let t = applist (mkLambda @@ -646,7 +657,7 @@ let clever_rewrite_base_poly typ p result theorem gl = [| typ; result; mkRel 2; mkRel 1; occ; theorem |]))), [abstracted]) in - exact (EConstr.of_constr (applist(t,[mkNewMeta()]))) gl + exact (applist(t,[mkNewMeta()])) gl let clever_rewrite_base p result theorem gl = clever_rewrite_base_poly (Lazy.force coq_Z) p result theorem gl @@ -664,11 +675,10 @@ let clever_rewrite_gen_nat p result (t,args) = let clever_rewrite p vpath t gl = let full = pf_concl gl in - let full = EConstr.Unsafe.to_constr full in - let (abstracted,occ) = abstract_path (Lazy.force coq_Z) (List.rev p) full in - let vargs = List.map (fun p -> occurrence p occ) vpath in + let (abstracted,occ) = abstract_path (project gl) (Lazy.force coq_Z) (List.rev p) full in + let vargs = List.map (fun p -> occurrence (project gl) p occ) vpath in let t' = applist(t, (vargs @ [abstracted])) in - exact (EConstr.of_constr (applist(t',[mkNewMeta()]))) gl + exact (applist(t',[mkNewMeta()])) gl let rec shuffle p (t1,t2) = match t1,t2 with @@ -910,10 +920,10 @@ let rec negate p = function | Oz i -> [focused_simpl p],Oz(neg i) | Oufo c -> [], Oufo (mkApp (Lazy.force coq_Zopp, [| c |])) -let rec transform p t = +let rec transform sigma p t = let default isnat t' = try - let v,th,_ = find_constr t' in + let v,th,_ = find_constr sigma t' in [clever_rewrite_base p (mkVar v) (mkVar th)], Oatom v with e when CErrors.noncritical e -> let v = new_identifier_var () @@ -921,29 +931,29 @@ let rec transform p t = hide_constr t' v th isnat; [clever_rewrite_base p (mkVar v) (mkVar th)], Oatom v in - try match destructurate_term t with + try match destructurate_term sigma t with | Kapp(Zplus,[t1;t2]) -> - let tac1,t1' = transform (P_APP 1 :: p) t1 - and tac2,t2' = transform (P_APP 2 :: p) t2 in + let tac1,t1' = transform sigma (P_APP 1 :: p) t1 + and tac2,t2' = transform sigma (P_APP 2 :: p) t2 in let tac,t' = shuffle p (t1',t2') in tac1 @ tac2 @ tac, t' | Kapp(Zminus,[t1;t2]) -> let tac,t = - transform p + transform sigma p (mkApp (Lazy.force coq_Zplus, [| t1; (mkApp (Lazy.force coq_Zopp, [| t2 |])) |])) in Proofview.V82.of_tactic (unfold sp_Zminus) :: tac,t | Kapp(Zsucc,[t1]) -> - let tac,t = transform p (mkApp (Lazy.force coq_Zplus, + let tac,t = transform sigma p (mkApp (Lazy.force coq_Zplus, [| t1; mk_integer one |])) in Proofview.V82.of_tactic (unfold sp_Zsucc) :: tac,t | Kapp(Zpred,[t1]) -> - let tac,t = transform p (mkApp (Lazy.force coq_Zplus, + let tac,t = transform sigma p (mkApp (Lazy.force coq_Zplus, [| t1; mk_integer negone |])) in Proofview.V82.of_tactic (unfold sp_Zpred) :: tac,t | Kapp(Zmult,[t1;t2]) -> - let tac1,t1' = transform (P_APP 1 :: p) t1 - and tac2,t2' = transform (P_APP 2 :: p) t2 in + let tac1,t1' = transform sigma (P_APP 1 :: p) t1 + and tac2,t2' = transform sigma (P_APP 2 :: p) t2 in begin match t1',t2' with | (_,Oz n) -> let tac,t' = scalar p n t1' in tac1 @ tac2 @ tac,t' | (Oz n,_) -> @@ -954,11 +964,11 @@ let rec transform p t = | _ -> default false t end | Kapp((Zpos|Zneg|Z0),_) -> - (try ([],Oz(recognize_number t)) + (try ([],Oz(recognize_number sigma t)) with e when CErrors.noncritical e -> default false t) | Kvar s -> [],Oatom s | Kapp(Zopp,[t]) -> - let tac,t' = transform (P_APP 1 :: p) t in + let tac,t' = transform sigma (P_APP 1 :: p) t in let tac',t'' = negate p t' in tac @ tac', t'' | Kapp(Z_of_nat,[t']) -> default true t' @@ -1088,13 +1098,12 @@ let replay_history tactic_normalisation = let p_initial = [P_APP 2;P_TYPE] in let tac = shuffle_cancel p_initial e1.body in let solve_le = - let not_sup_sup = mkApp (Universes.constr_of_global (build_coq_eq ()), + let not_sup_sup = mkApp (Lazy.force coq_eq, [| Lazy.force coq_comparison; Lazy.force coq_Gt; Lazy.force coq_Gt |]) in - let not_sup_sup = EConstr.of_constr not_sup_sup in Tacticals.New.tclTHENS (Tacticals.New.tclTHENLIST [ unfold sp_Zle; @@ -1120,7 +1129,7 @@ let replay_history tactic_normalisation = let state_eg = mk_eq eq1 rhs in let tac = scalar_norm_add [P_APP 3] e2.body in Tacticals.New.tclTHENS - (cut (EConstr.of_constr state_eg)) + (cut state_eg) [ Tacticals.New.tclTHENS (Tacticals.New.tclTHENLIST [ (intros_using [aux]); @@ -1189,7 +1198,7 @@ let replay_history tactic_normalisation = if e1.kind == DISE then let tac = scalar_norm [P_APP 3] e2.body in Tacticals.New.tclTHENS - (cut (EConstr.of_constr state_eq)) + (cut state_eq) [Tacticals.New.tclTHENLIST [ (intros_using [aux1]); (generalize_tac @@ -1201,7 +1210,7 @@ let replay_history tactic_normalisation = Tacticals.New.tclTHEN (Proofview.V82.tactic (mk_then tac)) reflexivity ] else let tac = scalar_norm [P_APP 3] e2.body in - Tacticals.New.tclTHENS (cut (EConstr.of_constr state_eq)) + Tacticals.New.tclTHENS (cut state_eq) [ Tacticals.New.tclTHENS (cut (mk_gt kk izero)) @@ -1231,7 +1240,7 @@ let replay_history tactic_normalisation = scalar_norm [P_APP 3] e1.body in Tacticals.New.tclTHENS - (cut (EConstr.of_constr (mk_eq eq1 (mk_inv eq2)))) + (cut (mk_eq eq1 (mk_inv eq2))) [Tacticals.New.tclTHENLIST [ (intros_using [aux]); (generalize_tac [mkApp (Lazy.force coq_OMEGA8, @@ -1249,7 +1258,7 @@ let replay_history tactic_normalisation = and eq2 = val_of(decompile orig) in let vid = unintern_id v in let theorem = - mkApp (build_coq_ex (), [| + mkApp (Lazy.force coq_ex, [| Lazy.force coq_Z; mkLambda (Name vid, @@ -1264,7 +1273,7 @@ let replay_history tactic_normalisation = shuffle_mult_right p_initial orig.body m ({c= negone;v= v}::def.body) in Tacticals.New.tclTHENS - (cut (EConstr.of_constr theorem)) + (cut theorem) [Tacticals.New.tclTHENLIST [ (intros_using [aux]); (elim_id aux); @@ -1277,7 +1286,7 @@ let replay_history tactic_normalisation = (clear [aux]); (intros_using [id]); (loop l) ]; - Tacticals.New.tclTHEN (exists_tac (EConstr.of_constr eq1)) reflexivity ] + Tacticals.New.tclTHEN (exists_tac eq1) reflexivity ] | SPLIT_INEQ(e,(e1,act1),(e2,act2)) :: l -> let id1 = new_identifier () and id2 = new_identifier () in @@ -1287,7 +1296,7 @@ let replay_history tactic_normalisation = let tac2 = scalar_norm_add [P_APP 2;P_TYPE] e.body in let eq = val_of(decompile e) in Tacticals.New.tclTHENS - (simplest_elim (EConstr.of_constr (applist (Lazy.force coq_OMEGA19, [eq; mkVar id])))) + (simplest_elim (applist (Lazy.force coq_OMEGA19, [eq; mkVar id]))) [Tacticals.New.tclTHENLIST [ Proofview.V82.tactic (mk_then tac1); (intros_using [id1]); (loop act1) ]; Tacticals.New.tclTHENLIST [ Proofview.V82.tactic (mk_then tac2); (intros_using [id2]); (loop act2) ]] | SUM(e3,(k1,e1),(k2,e2)) :: l -> @@ -1360,15 +1369,15 @@ let replay_history tactic_normalisation = in loop -let normalize p_initial t = - let (tac,t') = transform p_initial t in +let normalize sigma p_initial t = + let (tac,t') = transform sigma p_initial t in let (tac',t'') = condense p_initial t' in let (tac'',t''') = clear_zero p_initial t'' in tac @ tac' @ tac'' , t''' -let normalize_equation id flag theorem pos t t1 t2 (tactic,defs) = +let normalize_equation sigma id flag theorem pos t t1 t2 (tactic,defs) = let p_initial = [P_APP pos ;P_TYPE] in - let (tac,t') = normalize p_initial t in + let (tac,t') = normalize sigma p_initial t in let shift_left = tclTHEN (Proofview.V82.of_tactic (generalize_tac [mkApp (theorem, [| t1; t2; mkVar id |]) ])) @@ -1383,34 +1392,35 @@ let normalize_equation id flag theorem pos t t1 t2 (tactic,defs) = (tactic,defs) let destructure_omega gl tac_def (id,c) = + let sigma = project gl in if String.equal (atompart_of_id id) "State" then tac_def else - try match destructurate_prop c with + try match destructurate_prop sigma c with | Kapp(Eq,[typ;t1;t2]) - when begin match destructurate_type (EConstr.Unsafe.to_constr (pf_nf gl (EConstr.of_constr typ))) with Kapp(Z,[]) -> true | _ -> false end -> + when begin match destructurate_type sigma (pf_nf gl typ) with Kapp(Z,[]) -> true | _ -> false end -> let t = mk_plus t1 (mk_inv t2) in - normalize_equation + normalize_equation sigma id EQUA (Lazy.force coq_Zegal_left) 2 t t1 t2 tac_def | Kapp(Zne,[t1;t2]) -> let t = mk_plus t1 (mk_inv t2) in - normalize_equation + normalize_equation sigma id DISE (Lazy.force coq_Zne_left) 1 t t1 t2 tac_def | Kapp(Zle,[t1;t2]) -> let t = mk_plus t2 (mk_inv t1) in - normalize_equation + normalize_equation sigma id INEQ (Lazy.force coq_Zle_left) 2 t t1 t2 tac_def | Kapp(Zlt,[t1;t2]) -> let t = mk_plus (mk_plus t2 (mk_integer negone)) (mk_inv t1) in - normalize_equation + normalize_equation sigma id INEQ (Lazy.force coq_Zlt_left) 2 t t1 t2 tac_def | Kapp(Zge,[t1;t2]) -> let t = mk_plus t1 (mk_inv t2) in - normalize_equation + normalize_equation sigma id INEQ (Lazy.force coq_Zge_left) 2 t t1 t2 tac_def | Kapp(Zgt,[t1;t2]) -> let t = mk_plus (mk_plus t1 (mk_integer negone)) (mk_inv t2) in - normalize_equation + normalize_equation sigma id INEQ (Lazy.force coq_Zgt_left) 2 t t1 t2 tac_def | _ -> tac_def with e when catchable_exception e -> tac_def @@ -1426,7 +1436,6 @@ let coq_omega = Proofview.Goal.nf_enter { enter = begin fun gl -> clear_constr_tables (); let hyps_types = Tacmach.New.pf_hyps_types gl in - let hyps_types = List.map (on_snd EConstr.Unsafe.to_constr) hyps_types in let destructure_omega = Tacmach.New.of_old destructure_omega gl in let tactic_normalisation, system = List.fold_left destructure_omega ([],[]) hyps_types in @@ -1438,7 +1447,7 @@ let coq_omega = let i = new_id () in tag_hypothesis id i; (Tacticals.New.tclTHENLIST [ - (simplest_elim (EConstr.of_constr (applist (Lazy.force coq_intro_Z, [t])))); + (simplest_elim (applist (Lazy.force coq_intro_Z, [t]))); (intros_using [v; id]); (elim_id id); (clear [id]); @@ -1449,7 +1458,7 @@ let coq_omega = constant = zero; id = i} :: sys else (Tacticals.New.tclTHENLIST [ - (simplest_elim (EConstr.of_constr (applist (Lazy.force coq_new_var, [t])))); + (simplest_elim (applist (Lazy.force coq_new_var, [t]))); (intros_using [v;th]); tac ]), sys) @@ -1480,7 +1489,8 @@ let nat_inject = Proofview.Goal.nf_enter { enter = begin fun gl -> let is_conv = Tacmach.New.pf_apply Reductionops.is_conv gl in let rec explore p t : unit Proofview.tactic = - try match destructurate_term t with + Proofview.tclEVARMAP >>= fun sigma -> + try match destructurate_term sigma t with | Kapp(Plus,[t1;t2]) -> Tacticals.New.tclTHENLIST [ Proofview.V82.tactic (clever_rewrite_gen p (mk_plus (mk_inj t1) (mk_inj t2)) @@ -1499,7 +1509,7 @@ let nat_inject = let id = new_identifier () in Tacticals.New.tclTHENS (Tacticals.New.tclTHEN - (simplest_elim (EConstr.of_constr (applist (Lazy.force coq_le_gt_dec, [t2;t1])))) + (simplest_elim (applist (Lazy.force coq_le_gt_dec, [t2;t1]))) (intros_using [id])) [ Tacticals.New.tclTHENLIST [ @@ -1516,14 +1526,14 @@ let nat_inject = ] | Kapp(S,[t']) -> let rec is_number t = - try match destructurate_term t with + try match destructurate_term sigma t with Kapp(S,[t]) -> is_number t | Kapp(O,[]) -> true | _ -> false with e when catchable_exception e -> false in let rec loop p t : unit Proofview.tactic = - try match destructurate_term t with + try match destructurate_term sigma t with Kapp(S,[t]) -> (Tacticals.New.tclTHEN (Proofview.V82.tactic (clever_rewrite_gen p @@ -1549,7 +1559,8 @@ let nat_inject = and loop = function | [] -> Proofview.tclUNIT () | (i,t)::lit -> - begin try match destructurate_prop t with + Proofview.tclEVARMAP >>= fun sigma -> + begin try match destructurate_prop sigma t with Kapp(Le,[t1;t2]) -> Tacticals.New.tclTHENLIST [ (generalize_tac @@ -1596,7 +1607,7 @@ let nat_inject = (loop lit) ] | Kapp(Eq,[typ;t1;t2]) -> - if is_conv (EConstr.of_constr typ) (EConstr.of_constr (Lazy.force coq_nat)) then + if is_conv typ (Lazy.force coq_nat) then Tacticals.New.tclTHENLIST [ (generalize_tac [mkApp (Lazy.force coq_inj_eq, [| t1;t2;mkVar i |]) ]); @@ -1610,7 +1621,6 @@ let nat_inject = with e when catchable_exception e -> loop lit end in let hyps_types = Tacmach.New.pf_hyps_types gl in - let hyps_types = List.map (on_snd EConstr.Unsafe.to_constr) hyps_types in loop (List.rev hyps_types) end } @@ -1647,7 +1657,7 @@ let not_binop = function exception Undecidable let rec decidability gl t = - match destructurate_prop t with + match destructurate_prop (project gl) t with | Kapp(Or,[t1;t2]) -> mkApp (Lazy.force coq_dec_or, [| t1; t2; decidability gl t1; decidability gl t2 |]) @@ -1665,7 +1675,7 @@ let rec decidability gl t = | Kapp(Not,[t1]) -> mkApp (Lazy.force coq_dec_not, [| t1; decidability gl t1 |]) | Kapp(Eq,[typ;t1;t2]) -> - begin match destructurate_type (EConstr.Unsafe.to_constr (pf_nf gl (EConstr.of_constr typ))) with + begin match destructurate_type (project gl) (pf_nf gl typ) with | Kapp(Z,[]) -> mkApp (Lazy.force coq_dec_eq, [| t1;t2 |]) | Kapp(Nat,[]) -> mkApp (Lazy.force coq_dec_eq_nat, [| t1;t2 |]) | _ -> raise Undecidable @@ -1696,6 +1706,15 @@ let onClearedName2 id tac = Tacticals.New.tclTHENLIST [ introduction id1; introduction id2; tac id1 id2 ] end }) +let local_assum (na, t) = + let inj = EConstr.Unsafe.to_constr in + LocalAssum (na, inj t) + +let rec is_Prop sigma c = match EConstr.kind sigma c with + | Sort (Prop Null) -> true + | Cast (c,_,_) -> is_Prop sigma c + | _ -> false + let destructure_hyps = Proofview.Goal.nf_enter { enter = begin fun gl -> let type_of = Tacmach.New.pf_unsafe_type_of gl in @@ -1705,46 +1724,47 @@ let destructure_hyps = | [] -> (Tacticals.New.tclTHEN nat_inject coq_omega) | decl::lit -> let i = NamedDecl.get_id decl in - begin try match destructurate_prop (NamedDecl.get_type decl) with + Proofview.tclEVARMAP >>= fun sigma -> + begin try match destructurate_prop sigma (EConstr.of_constr (NamedDecl.get_type decl)) with | Kapp(False,[]) -> elim_id i | Kapp((Zle|Zge|Zgt|Zlt|Zne),[t1;t2]) -> loop lit | Kapp(Or,[t1;t2]) -> (Tacticals.New.tclTHENS (elim_id i) - [ onClearedName i (fun i -> (loop (LocalAssum (i,t1)::lit))); - onClearedName i (fun i -> (loop (LocalAssum (i,t2)::lit))) ]) + [ onClearedName i (fun i -> (loop (local_assum (i,t1)::lit))); + onClearedName i (fun i -> (loop (local_assum (i,t2)::lit))) ]) | Kapp(And,[t1;t2]) -> Tacticals.New.tclTHEN (elim_id i) (onClearedName2 i (fun i1 i2 -> - loop (LocalAssum (i1,t1) :: LocalAssum (i2,t2) :: lit))) + loop (local_assum (i1,t1) :: local_assum (i2,t2) :: lit))) | Kapp(Iff,[t1;t2]) -> Tacticals.New.tclTHEN (elim_id i) (onClearedName2 i (fun i1 i2 -> - loop (LocalAssum (i1,mkArrow t1 t2) :: LocalAssum (i2,mkArrow t2 t1) :: lit))) + loop (local_assum (i1,mkArrow t1 t2) :: local_assum (i2,mkArrow t2 t1) :: lit))) | Kimp(t1,t2) -> (* t1 and t2 might be in Type rather than Prop. For t1, the decidability check will ensure being Prop. *) - if is_Prop (EConstr.Unsafe.to_constr (type_of (EConstr.of_constr t2))) + if is_Prop sigma (type_of t2) then let d1 = decidability t1 in Tacticals.New.tclTHENLIST [ (generalize_tac [mkApp (Lazy.force coq_imp_simp, [| t1; t2; d1; mkVar i|])]); (onClearedName i (fun i -> - (loop (LocalAssum (i,mk_or (mk_not t1) t2) :: lit)))) + (loop (local_assum (i,mk_or (mk_not t1) t2) :: lit)))) ] else loop lit | Kapp(Not,[t]) -> - begin match destructurate_prop t with + begin match destructurate_prop sigma t with Kapp(Or,[t1;t2]) -> Tacticals.New.tclTHENLIST [ (generalize_tac [mkApp (Lazy.force coq_not_or,[| t1; t2; mkVar i |])]); (onClearedName i (fun i -> - (loop (LocalAssum (i,mk_and (mk_not t1) (mk_not t2)) :: lit)))) + (loop (local_assum (i,mk_and (mk_not t1) (mk_not t2)) :: lit)))) ] | Kapp(And,[t1;t2]) -> let d1 = decidability t1 in @@ -1753,7 +1773,7 @@ let destructure_hyps = [mkApp (Lazy.force coq_not_and, [| t1; t2; d1; mkVar i |])]); (onClearedName i (fun i -> - (loop (LocalAssum (i,mk_or (mk_not t1) (mk_not t2)) :: lit)))) + (loop (local_assum (i,mk_or (mk_not t1) (mk_not t2)) :: lit)))) ] | Kapp(Iff,[t1;t2]) -> let d1 = decidability t1 in @@ -1763,7 +1783,7 @@ let destructure_hyps = [mkApp (Lazy.force coq_not_iff, [| t1; t2; d1; d2; mkVar i |])]); (onClearedName i (fun i -> - (loop (LocalAssum (i, mk_or (mk_and t1 (mk_not t2)) + (loop (local_assum (i, mk_or (mk_and t1 (mk_not t2)) (mk_and (mk_not t1) t2)) :: lit)))) ] | Kimp(t1,t2) -> @@ -1775,14 +1795,14 @@ let destructure_hyps = [mkApp (Lazy.force coq_not_imp, [| t1; t2; d1; mkVar i |])]); (onClearedName i (fun i -> - (loop (LocalAssum (i,mk_and t1 (mk_not t2)) :: lit)))) + (loop (local_assum (i,mk_and t1 (mk_not t2)) :: lit)))) ] | Kapp(Not,[t]) -> let d = decidability t in Tacticals.New.tclTHENLIST [ (generalize_tac [mkApp (Lazy.force coq_not_not, [| t; d; mkVar i |])]); - (onClearedName i (fun i -> (loop (LocalAssum (i,t) :: lit)))) + (onClearedName i (fun i -> (loop (local_assum (i,t) :: lit)))) ] | Kapp(op,[t1;t2]) -> (try @@ -1795,32 +1815,32 @@ let destructure_hyps = with Not_found -> loop lit) | Kapp(Eq,[typ;t1;t2]) -> if !old_style_flag then begin - match destructurate_type (EConstr.Unsafe.to_constr (pf_nf (EConstr.of_constr typ))) with + match destructurate_type sigma (pf_nf typ) with | Kapp(Nat,_) -> Tacticals.New.tclTHENLIST [ (simplest_elim - (EConstr.of_constr (mkApp - (Lazy.force coq_not_eq, [|t1;t2;mkVar i|])))); + (mkApp + (Lazy.force coq_not_eq, [|t1;t2;mkVar i|]))); (onClearedName i (fun _ -> loop lit)) ] | Kapp(Z,_) -> Tacticals.New.tclTHENLIST [ (simplest_elim - (EConstr.of_constr (mkApp - (Lazy.force coq_not_Zeq, [|t1;t2;mkVar i|])))); + (mkApp + (Lazy.force coq_not_Zeq, [|t1;t2;mkVar i|]))); (onClearedName i (fun _ -> loop lit)) ] | _ -> loop lit end else begin - match destructurate_type (EConstr.Unsafe.to_constr (pf_nf (EConstr.of_constr typ))) with + match destructurate_type sigma (pf_nf typ) with | Kapp(Nat,_) -> (Tacticals.New.tclTHEN - (convert_hyp_no_check (NamedDecl.set_type (mkApp (Lazy.force coq_neq, [| t1;t2|])) + (convert_hyp_no_check (NamedDecl.set_type (EConstr.Unsafe.to_constr (mkApp (Lazy.force coq_neq, [| t1;t2|]))) decl)) (loop lit)) | Kapp(Z,_) -> (Tacticals.New.tclTHEN - (convert_hyp_no_check (NamedDecl.set_type (mkApp (Lazy.force coq_Zne, [| t1;t2|])) + (convert_hyp_no_check (NamedDecl.set_type (EConstr.Unsafe.to_constr (mkApp (Lazy.force coq_Zne, [| t1;t2|]))) decl)) (loop lit)) | _ -> loop lit @@ -1842,7 +1862,8 @@ let destructure_goal = let concl = Proofview.Goal.concl gl in let decidability = Tacmach.New.of_old decidability gl in let rec loop t = - match destructurate_prop t with + Proofview.tclEVARMAP >>= fun sigma -> + match destructurate_prop sigma t with | Kapp(Not,[t]) -> (Tacticals.New.tclTHEN (Tacticals.New.tclTHEN (unfold sp_not) intro) @@ -1855,13 +1876,12 @@ let destructure_goal = let dec = decidability t in Tacticals.New.tclTHEN (Proofview.V82.tactic (Tacmach.refine - (EConstr.of_constr (mkApp (Lazy.force coq_dec_not_not, [| t; dec; mkNewMeta () |]))))) + (mkApp (Lazy.force coq_dec_not_not, [| t; dec; mkNewMeta () |])))) intro - with Undecidable -> Tactics.elim_type (EConstr.of_constr (build_coq_False ())) + with Undecidable -> Tactics.elim_type (Lazy.force coq_False) in Tacticals.New.tclTHEN goal_tac destructure_hyps in - let concl = EConstr.Unsafe.to_constr concl in (loop concl) end } -- cgit v1.2.3 From 8beca748d992cd08e2dd7448c8b28dadbcea4a16 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sat, 26 Nov 2016 01:09:11 +0100 Subject: Cleaning up interfaces. We make mli files look to what they were looking before the move to EConstr by opening this module. --- pretyping/find_subterm.ml | 4 +- pretyping/find_subterm.mli | 15 ++++---- pretyping/pretype_errors.mli | 87 ++++++++++++++++++++++---------------------- pretyping/pretyping.ml | 2 +- pretyping/pretyping.mli | 31 ++++++++-------- pretyping/tacred.ml | 1 - pretyping/tacred.mli | 17 +++++---- pretyping/unification.ml | 3 +- tactics/leminv.mli | 3 +- tactics/tacticals.mli | 31 ++++++++-------- tactics/tactics.ml | 1 - 11 files changed, 99 insertions(+), 96 deletions(-) diff --git a/pretyping/find_subterm.ml b/pretyping/find_subterm.ml index 15409f2b8..d09686f6e 100644 --- a/pretyping/find_subterm.ml +++ b/pretyping/find_subterm.ml @@ -141,8 +141,8 @@ let replace_term_occ_gen_modulo sigma occs like_first test bywhat cl occ t = let replace_term_occ_modulo evd occs test bywhat t = let occs',like_first = match occs with AtOccs occs -> occs,false | LikeFirst -> AllOccurrences,true in - EConstr.Unsafe.to_constr (proceed_with_occurrences - (replace_term_occ_gen_modulo evd occs' like_first test bywhat None) occs' t) + proceed_with_occurrences + (replace_term_occ_gen_modulo evd occs' like_first test bywhat None) occs' t let replace_term_occ_decl_modulo evd occs test bywhat d = let (plocs,hyploc),like_first = diff --git a/pretyping/find_subterm.mli b/pretyping/find_subterm.mli index c7db84e3c..3d2ebb72d 100644 --- a/pretyping/find_subterm.mli +++ b/pretyping/find_subterm.mli @@ -11,11 +11,12 @@ open Term open Evd open Pretype_errors open Environ +open EConstr (** Finding subterms, possibly up to some unification function, possibly at some given occurrences *) -exception NotUnifiable of (EConstr.constr * EConstr.constr * unification_error) option +exception NotUnifiable of (constr * constr * unification_error) option exception SubtermUnificationError of subterm_unification_error @@ -26,7 +27,7 @@ exception SubtermUnificationError of subterm_unification_error with None. *) type 'a testing_function = { - match_fun : 'a -> EConstr.constr -> 'a; + match_fun : 'a -> constr -> 'a; merge_fun : 'a -> 'a -> 'a; mutable testing_state : 'a; mutable last_found : position_reporting option @@ -34,7 +35,7 @@ type 'a testing_function = { (** This is the basic testing function, looking for exact matches of a closed term *) -val make_eq_univs_test : env -> evar_map -> EConstr.constr -> evar_map testing_function +val make_eq_univs_test : env -> evar_map -> constr -> evar_map testing_function (** [replace_term_occ_modulo occl test mk c] looks in [c] for subterm modulo a testing function [test] and replaces successfully @@ -42,27 +43,27 @@ val make_eq_univs_test : env -> evar_map -> EConstr.constr -> evar_map testing_f ()]; it turns a NotUnifiable exception raised by the testing function into a SubtermUnificationError. *) val replace_term_occ_modulo : evar_map -> occurrences or_like_first -> - 'a testing_function -> (unit -> EConstr.constr) -> EConstr.constr -> constr + 'a testing_function -> (unit -> constr) -> constr -> constr (** [replace_term_occ_decl_modulo] is similar to [replace_term_occ_modulo] but for a named_declaration. *) val replace_term_occ_decl_modulo : evar_map -> (occurrences * hyp_location_flag) or_like_first -> - 'a testing_function -> (unit -> EConstr.constr) -> + 'a testing_function -> (unit -> constr) -> Context.Named.Declaration.t -> Context.Named.Declaration.t (** [subst_closed_term_occ occl c d] replaces occurrences of closed [c] at positions [occl] by [Rel 1] in [d] (see also Note OCC), unifying universes which results in a set of constraints. *) val subst_closed_term_occ : env -> evar_map -> occurrences or_like_first -> - EConstr.constr -> EConstr.constr -> constr * evar_map + constr -> constr -> constr * evar_map (** [subst_closed_term_occ_decl evd occl c decl] replaces occurrences of closed [c] at positions [occl] by [Rel 1] in [decl]. *) val subst_closed_term_occ_decl : env -> evar_map -> (occurrences * hyp_location_flag) or_like_first -> - EConstr.constr -> Context.Named.Declaration.t -> Context.Named.Declaration.t * evar_map + constr -> Context.Named.Declaration.t -> Context.Named.Declaration.t * evar_map (** Miscellaneous *) val error_invalid_occurrence : int list -> 'a diff --git a/pretyping/pretype_errors.mli b/pretyping/pretype_errors.mli index 0ebe4817c..7cef14339 100644 --- a/pretyping/pretype_errors.mli +++ b/pretyping/pretype_errors.mli @@ -9,52 +9,53 @@ open Names open Term open Environ +open EConstr open Type_errors (** {6 The type of errors raised by the pretyper } *) type unification_error = - | OccurCheck of existential_key * EConstr.constr - | NotClean of EConstr.existential * env * EConstr.constr + | OccurCheck of existential_key * constr + | NotClean of existential * env * constr | NotSameArgSize | NotSameHead | NoCanonicalStructure - | ConversionFailed of env * EConstr.constr * EConstr.constr + | ConversionFailed of env * constr * constr | MetaOccurInBody of existential_key - | InstanceNotSameType of existential_key * env * EConstr.types * EConstr.types + | InstanceNotSameType of existential_key * env * types * types | UnifUnivInconsistency of Univ.univ_inconsistency | CannotSolveConstraint of Evd.evar_constraint * unification_error | ProblemBeyondCapabilities type position = (Id.t * Locus.hyp_location_flag) option -type position_reporting = (position * int) * EConstr.t +type position_reporting = (position * int) * constr -type subterm_unification_error = bool * position_reporting * position_reporting * (EConstr.constr * EConstr.constr * unification_error) option +type subterm_unification_error = bool * position_reporting * position_reporting * (constr * constr * unification_error) option -type type_error = (EConstr.constr, EConstr.types) ptype_error +type type_error = (constr, types) ptype_error type pretype_error = (** Old Case *) - | CantFindCaseType of EConstr.constr + | CantFindCaseType of constr (** Type inference unification *) - | ActualTypeNotCoercible of EConstr.unsafe_judgment * EConstr.types * unification_error + | ActualTypeNotCoercible of unsafe_judgment * types * unification_error (** Tactic Unification *) - | UnifOccurCheck of existential_key * EConstr.constr + | UnifOccurCheck of existential_key * constr | UnsolvableImplicit of existential_key * Evd.unsolvability_explanation option - | CannotUnify of EConstr.constr * EConstr.constr * unification_error option - | CannotUnifyLocal of EConstr.constr * EConstr.constr * EConstr.constr - | CannotUnifyBindingType of constr * constr - | CannotGeneralize of constr - | NoOccurrenceFound of EConstr.constr * Id.t option - | CannotFindWellTypedAbstraction of EConstr.constr * EConstr.constr list * (env * type_error) option - | WrongAbstractionType of Name.t * EConstr.constr * EConstr.types * EConstr.types + | CannotUnify of constr * constr * unification_error option + | CannotUnifyLocal of constr * constr * constr + | CannotUnifyBindingType of Constr.constr * Constr.constr + | CannotGeneralize of Constr.constr + | NoOccurrenceFound of constr * Id.t option + | CannotFindWellTypedAbstraction of constr * constr list * (env * type_error) option + | WrongAbstractionType of Name.t * constr * types * types | AbstractionOverMeta of Name.t * Name.t - | NonLinearUnification of Name.t * EConstr.constr + | NonLinearUnification of Name.t * constr (** Pretyping *) | VarNotFound of Id.t - | UnexpectedType of EConstr.constr * EConstr.constr - | NotProduct of EConstr.constr + | UnexpectedType of constr * constr + | NotProduct of constr | TypingError of type_error | CannotUnifyOccurrences of subterm_unification_error | UnsatisfiableConstraints of @@ -67,85 +68,85 @@ val precatchable_exception : exn -> bool (** Raising errors *) val error_actual_type : - ?loc:Loc.t -> env -> Evd.evar_map -> EConstr.unsafe_judgment -> EConstr.constr -> + ?loc:Loc.t -> env -> Evd.evar_map -> unsafe_judgment -> constr -> unification_error -> 'b val error_actual_type_core : - ?loc:Loc.t -> env -> Evd.evar_map -> EConstr.unsafe_judgment -> EConstr.constr -> 'b + ?loc:Loc.t -> env -> Evd.evar_map -> unsafe_judgment -> constr -> 'b val error_cant_apply_not_functional : ?loc:Loc.t -> env -> Evd.evar_map -> - EConstr.unsafe_judgment -> EConstr.unsafe_judgment array -> 'b + unsafe_judgment -> unsafe_judgment array -> 'b val error_cant_apply_bad_type : - ?loc:Loc.t -> env -> Evd.evar_map -> int * EConstr.constr * EConstr.constr -> - EConstr.unsafe_judgment -> EConstr.unsafe_judgment array -> 'b + ?loc:Loc.t -> env -> Evd.evar_map -> int * constr * constr -> + unsafe_judgment -> unsafe_judgment array -> 'b val error_case_not_inductive : - ?loc:Loc.t -> env -> Evd.evar_map -> EConstr.unsafe_judgment -> 'b + ?loc:Loc.t -> env -> Evd.evar_map -> unsafe_judgment -> 'b val error_ill_formed_branch : ?loc:Loc.t -> env -> Evd.evar_map -> - EConstr.constr -> pconstructor -> EConstr.constr -> EConstr.constr -> 'b + constr -> pconstructor -> constr -> constr -> 'b val error_number_branches : ?loc:Loc.t -> env -> Evd.evar_map -> - EConstr.unsafe_judgment -> int -> 'b + unsafe_judgment -> int -> 'b val error_ill_typed_rec_body : ?loc:Loc.t -> env -> Evd.evar_map -> - int -> Name.t array -> EConstr.unsafe_judgment array -> EConstr.types array -> 'b + int -> Name.t array -> unsafe_judgment array -> types array -> 'b val error_elim_arity : ?loc:Loc.t -> env -> Evd.evar_map -> - pinductive -> sorts_family list -> EConstr.constr -> - EConstr.unsafe_judgment -> (sorts_family * sorts_family * arity_error) option -> 'b + pinductive -> sorts_family list -> constr -> + unsafe_judgment -> (sorts_family * sorts_family * arity_error) option -> 'b val error_not_a_type : - ?loc:Loc.t -> env -> Evd.evar_map -> EConstr.unsafe_judgment -> 'b + ?loc:Loc.t -> env -> Evd.evar_map -> unsafe_judgment -> 'b val error_assumption : - ?loc:Loc.t -> env -> Evd.evar_map -> EConstr.unsafe_judgment -> 'b + ?loc:Loc.t -> env -> Evd.evar_map -> unsafe_judgment -> 'b -val error_cannot_coerce : env -> Evd.evar_map -> EConstr.constr * EConstr.constr -> 'b +val error_cannot_coerce : env -> Evd.evar_map -> constr * constr -> 'b (** {6 Implicit arguments synthesis errors } *) -val error_occur_check : env -> Evd.evar_map -> existential_key -> EConstr.constr -> 'b +val error_occur_check : env -> Evd.evar_map -> existential_key -> constr -> 'b val error_unsolvable_implicit : ?loc:Loc.t -> env -> Evd.evar_map -> existential_key -> Evd.unsolvability_explanation option -> 'b val error_cannot_unify : ?loc:Loc.t -> env -> Evd.evar_map -> - ?reason:unification_error -> EConstr.constr * EConstr.constr -> 'b + ?reason:unification_error -> constr * constr -> 'b -val error_cannot_unify_local : env -> Evd.evar_map -> EConstr.constr * EConstr.constr * EConstr.constr -> 'b +val error_cannot_unify_local : env -> Evd.evar_map -> constr * constr * constr -> 'b val error_cannot_find_well_typed_abstraction : env -> Evd.evar_map -> - EConstr.constr -> EConstr.constr list -> (env * type_error) option -> 'b + constr -> constr list -> (env * type_error) option -> 'b val error_wrong_abstraction_type : env -> Evd.evar_map -> - Name.t -> EConstr.constr -> EConstr.types -> EConstr.types -> 'b + Name.t -> constr -> types -> types -> 'b val error_abstraction_over_meta : env -> Evd.evar_map -> metavariable -> metavariable -> 'b val error_non_linear_unification : env -> Evd.evar_map -> - metavariable -> EConstr.constr -> 'b + metavariable -> constr -> 'b (** {6 Ml Case errors } *) val error_cant_find_case_type : - ?loc:Loc.t -> env -> Evd.evar_map -> EConstr.constr -> 'b + ?loc:Loc.t -> env -> Evd.evar_map -> constr -> 'b (** {6 Pretyping errors } *) val error_unexpected_type : - ?loc:Loc.t -> env -> Evd.evar_map -> EConstr.constr -> EConstr.constr -> 'b + ?loc:Loc.t -> env -> Evd.evar_map -> constr -> constr -> 'b val error_not_product : - ?loc:Loc.t -> env -> Evd.evar_map -> EConstr.constr -> 'b + ?loc:Loc.t -> env -> Evd.evar_map -> constr -> 'b (** {6 Error in conversion from AST to glob_constr } *) diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml index f76f608d0..6b6800ac6 100644 --- a/pretyping/pretyping.ml +++ b/pretyping/pretyping.ml @@ -61,7 +61,7 @@ type ltac_var_map = { ltac_genargs : unbound_ltac_var_map; } type glob_constr_ltac_closure = ltac_var_map * glob_constr -type pure_open_constr = evar_map * Constr.constr +type pure_open_constr = evar_map * EConstr.constr (************************************************************************) (* This concerns Cases *) diff --git a/pretyping/pretyping.mli b/pretyping/pretyping.mli index 825d73f1f..47ad93d7e 100644 --- a/pretyping/pretyping.mli +++ b/pretyping/pretyping.mli @@ -16,6 +16,7 @@ open Names open Term open Environ open Evd +open EConstr open Glob_term open Evarutil open Misctypes @@ -25,7 +26,7 @@ open Misctypes val search_guard : Loc.t -> env -> int list list -> rec_declaration -> int array -type typing_constraint = OfType of EConstr.types | IsType | WithoutTypeConstraint +type typing_constraint = OfType of types | IsType | WithoutTypeConstraint type var_map = Pattern.constr_under_binders Id.Map.t type uconstr_var_map = Glob_term.closed_glob_constr Id.Map.t @@ -47,7 +48,7 @@ val empty_lvar : ltac_var_map type glob_constr_ltac_closure = ltac_var_map * glob_constr type pure_open_constr = evar_map * constr -type inference_hook = env -> evar_map -> evar -> evar_map * EConstr.constr +type inference_hook = env -> evar_map -> evar -> evar_map * constr type inference_flags = { use_typeclasses : bool; @@ -76,10 +77,10 @@ val allow_anonymous_refs : bool ref evar_map is modified explicitly or by side-effect. *) val understand_tcc : ?flags:inference_flags -> env -> evar_map -> - ?expected_type:typing_constraint -> glob_constr -> evar_map * EConstr.constr + ?expected_type:typing_constraint -> glob_constr -> evar_map * constr val understand_tcc_evars : ?flags:inference_flags -> env -> evar_map ref -> - ?expected_type:typing_constraint -> glob_constr -> EConstr.constr + ?expected_type:typing_constraint -> glob_constr -> constr (** More general entry point with evars from ltac *) @@ -95,7 +96,7 @@ val understand_tcc_evars : ?flags:inference_flags -> env -> evar_map ref -> val understand_ltac : inference_flags -> env -> evar_map -> ltac_var_map -> - typing_constraint -> glob_constr -> evar_map * EConstr.constr + typing_constraint -> glob_constr -> pure_open_constr (** Standard call to get a constr from a glob_constr, resolving implicit arguments and coercions, and compiling pattern-matching; @@ -105,21 +106,21 @@ val understand_ltac : inference_flags -> unresolved evar remains, expanding evars. *) val understand : ?flags:inference_flags -> ?expected_type:typing_constraint -> - env -> evar_map -> glob_constr -> constr Evd.in_evar_universe_context + env -> evar_map -> glob_constr -> Constr.constr Evd.in_evar_universe_context (** Idem but returns the judgment of the understood term *) val understand_judgment : env -> evar_map -> - glob_constr -> EConstr.unsafe_judgment Evd.in_evar_universe_context + glob_constr -> unsafe_judgment Evd.in_evar_universe_context (** Idem but do not fail on unresolved evars (type cl*) val understand_judgment_tcc : env -> evar_map ref -> - glob_constr -> EConstr.unsafe_judgment + glob_constr -> unsafe_judgment val type_uconstr : ?flags:inference_flags -> ?expected_type:typing_constraint -> - Geninterp.interp_sign -> Glob_term.closed_glob_constr -> EConstr.constr Tactypes.delayed_open + Geninterp.interp_sign -> Glob_term.closed_glob_constr -> constr Tactypes.delayed_open (** Trying to solve remaining evars and remaining conversion problems possibly using type classes, heuristics, external tactic solver @@ -139,21 +140,21 @@ val check_evars_are_solved : (** [check_evars env initial_sigma extended_sigma c] fails if some new unresolved evar remains in [c] *) -val check_evars : env -> evar_map -> evar_map -> EConstr.constr -> unit +val check_evars : env -> evar_map -> evar_map -> constr -> unit (**/**) (** Internal of Pretyping... *) val pretype : int -> bool -> type_constraint -> env -> evar_map ref -> - ltac_var_map -> glob_constr -> EConstr.unsafe_judgment + ltac_var_map -> glob_constr -> unsafe_judgment val pretype_type : int -> bool -> val_constraint -> env -> evar_map ref -> - ltac_var_map -> glob_constr -> EConstr.unsafe_type_judgment + ltac_var_map -> glob_constr -> unsafe_type_judgment val ise_pretype_gen : inference_flags -> env -> evar_map -> - ltac_var_map -> typing_constraint -> glob_constr -> evar_map * EConstr.constr + ltac_var_map -> typing_constraint -> glob_constr -> evar_map * constr (**/**) @@ -163,5 +164,5 @@ val interp_sort : ?loc:Loc.t -> evar_map -> glob_sort -> evar_map * sorts val interp_elimination_sort : glob_sort -> sorts_family val genarg_interp_hook : - (EConstr.types -> env -> evar_map -> unbound_ltac_var_map -> - Genarg.glob_generic_argument -> EConstr.constr * evar_map) Hook.t + (types -> env -> evar_map -> unbound_ltac_var_map -> + Genarg.glob_generic_argument -> constr * evar_map) Hook.t diff --git a/pretyping/tacred.ml b/pretyping/tacred.ml index 2b496f926..4abfc26fc 100644 --- a/pretyping/tacred.ml +++ b/pretyping/tacred.ml @@ -1154,7 +1154,6 @@ let abstract_scheme env sigma (locc,a) (c, sigma) = mkLambda (na,ta,c), sigma else let c', sigma' = subst_closed_term_occ env sigma (AtOccs locc) a c in - let c' = EConstr.of_constr c' in mkLambda (na,ta,c'), sigma' let pattern_occs loccs_trm = { e_redfun = begin fun env sigma c -> diff --git a/pretyping/tacred.mli b/pretyping/tacred.mli index 15b4e990d..a4499015d 100644 --- a/pretyping/tacred.mli +++ b/pretyping/tacred.mli @@ -10,6 +10,7 @@ open Names open Term open Environ open Evd +open EConstr open Reductionops open Pattern open Globnames @@ -17,7 +18,7 @@ open Locus open Univ type reduction_tactic_error = - InvalidAbstraction of env * evar_map * EConstr.constr * (env * Type_errors.type_error) + InvalidAbstraction of env * evar_map * constr * (env * Type_errors.type_error) exception ReductionTacticError of reduction_tactic_error @@ -58,10 +59,10 @@ val unfoldn : (occurrences * evaluable_global_reference) list -> reduction_function (** Fold *) -val fold_commands : EConstr.constr list -> reduction_function +val fold_commands : constr list -> reduction_function (** Pattern *) -val pattern_occs : (occurrences * EConstr.constr) list -> e_reduction_function +val pattern_occs : (occurrences * constr) list -> e_reduction_function (** Rem: Lazy strategies are defined in Reduction *) @@ -75,23 +76,23 @@ val cbv_norm_flags : CClosure.RedFlags.reds -> reduction_function (** [reduce_to_atomic_ind env sigma t] puts [t] in the form [t'=(I args)] with [I] an inductive definition; returns [I] and [t'] or fails with a user error *) -val reduce_to_atomic_ind : env -> evar_map -> EConstr.types -> pinductive * EConstr.types +val reduce_to_atomic_ind : env -> evar_map -> types -> pinductive * types (** [reduce_to_quantified_ind env sigma t] puts [t] in the form [t'=(x1:A1)..(xn:An)(I args)] with [I] an inductive definition; returns [I] and [t'] or fails with a user error *) -val reduce_to_quantified_ind : env -> evar_map -> EConstr.types -> pinductive * EConstr.types +val reduce_to_quantified_ind : env -> evar_map -> types -> pinductive * types (** [reduce_to_quantified_ref env sigma ref t] try to put [t] in the form [t'=(x1:A1)..(xn:An)(ref args)] and fails with user error if not possible *) val reduce_to_quantified_ref : - env -> evar_map -> global_reference -> EConstr.types -> EConstr.types + env -> evar_map -> global_reference -> types -> types val reduce_to_atomic_ref : - env -> evar_map -> global_reference -> EConstr.types -> EConstr.types + env -> evar_map -> global_reference -> types -> types val find_hnf_rectype : - env -> evar_map -> EConstr.types -> pinductive * EConstr.constr list + env -> evar_map -> types -> pinductive * constr list val contextually : bool -> occurrences * constr_pattern -> (patvar_map -> reduction_function) -> reduction_function diff --git a/pretyping/unification.ml b/pretyping/unification.ml index 5bb865310..20f27a15a 100644 --- a/pretyping/unification.ml +++ b/pretyping/unification.ml @@ -100,7 +100,6 @@ let abstract_scheme env evd c l lname_typ = if occur_meta evd a then mkLambda_name env (na,ta,t), evd else let t', evd' = Find_subterm.subst_closed_term_occ env evd locc a t in - let t' = EConstr.of_constr t' in mkLambda_name env (na,ta,t'), evd') (c,evd) (List.rev l) @@ -1656,7 +1655,7 @@ let make_abstraction_core name (test,out) env sigma c ty occs check_occs concl = | NoOccurrences -> concl | occ -> let occ = if likefirst then LikeFirst else AtOccs occ in - EConstr.of_constr (replace_term_occ_modulo sigma occ test mkvarid concl) + replace_term_occ_modulo sigma occ test mkvarid concl in let lastlhyp = if List.is_empty depdecls then None else Some (NamedDecl.get_id (List.last depdecls)) in diff --git a/tactics/leminv.mli b/tactics/leminv.mli index 58b82002d..26d4ac994 100644 --- a/tactics/leminv.mli +++ b/tactics/leminv.mli @@ -8,11 +8,12 @@ open Names open Term +open EConstr open Constrexpr open Misctypes val lemInv_clause : - quantified_hypothesis -> EConstr.constr -> Id.t list -> unit Proofview.tactic + quantified_hypothesis -> constr -> Id.t list -> unit Proofview.tactic val add_inversion_lemma_exn : Id.t -> constr_expr -> glob_sort -> bool -> (Id.t -> unit Proofview.tactic) -> diff --git a/tactics/tacticals.mli b/tactics/tacticals.mli index e4f110722..ba5452e33 100644 --- a/tactics/tacticals.mli +++ b/tactics/tacticals.mli @@ -9,6 +9,7 @@ open Pp open Names open Term +open EConstr open Tacmach open Proof_type open Locus @@ -58,25 +59,25 @@ val tclIFTHENTRYELSEMUST : tactic -> tactic -> tactic (** {6 Tacticals applying to hypotheses } *) val onNthHypId : int -> (Id.t -> tactic) -> tactic -val onNthHyp : int -> (EConstr.constr -> tactic) -> tactic +val onNthHyp : int -> (constr -> tactic) -> tactic val onNthDecl : int -> (Context.Named.Declaration.t -> tactic) -> tactic val onLastHypId : (Id.t -> tactic) -> tactic -val onLastHyp : (EConstr.constr -> tactic) -> tactic +val onLastHyp : (constr -> tactic) -> tactic val onLastDecl : (Context.Named.Declaration.t -> tactic) -> tactic val onNLastHypsId : int -> (Id.t list -> tactic) -> tactic -val onNLastHyps : int -> (constr list -> tactic) -> tactic +val onNLastHyps : int -> (Constr.constr list -> tactic) -> tactic val onNLastDecls : int -> (Context.Named.t -> tactic) -> tactic val lastHypId : goal sigma -> Id.t -val lastHyp : goal sigma -> EConstr.constr +val lastHyp : goal sigma -> constr val lastDecl : goal sigma -> Context.Named.Declaration.t val nLastHypsId : int -> goal sigma -> Id.t list -val nLastHyps : int -> goal sigma -> constr list +val nLastHyps : int -> goal sigma -> Constr.constr list val nLastDecls : int -> goal sigma -> Context.Named.t val afterHyp : Id.t -> goal sigma -> Context.Named.t -val ifOnHyp : (Id.t * EConstr.types -> bool) -> +val ifOnHyp : (Id.t * types -> bool) -> (Id.t -> tactic) -> (Id.t -> tactic) -> Id.t -> tactic @@ -99,9 +100,9 @@ val onClauseLR : (Id.t option -> tactic) -> clause -> tactic type branch_args = private { ity : pinductive; (** the type we were eliminating on *) - largs : EConstr.constr list; (** its arguments *) + largs : constr list; (** its arguments *) branchnum : int; (** the branch number *) - pred : EConstr.constr; (** the predicate we used *) + pred : constr; (** the predicate we used *) nassums : int; (** number of assumptions/letin to be introduced *) branchsign : bool list; (** the signature of the branch. true=assumption, false=let-in *) @@ -134,7 +135,7 @@ val elimination_sort_of_hyp : Id.t -> goal sigma -> sorts_family val elimination_sort_of_clause : Id.t option -> goal sigma -> sorts_family val pf_with_evars : (goal sigma -> Evd.evar_map * 'a) -> ('a -> tactic) -> tactic -val pf_constr_of_global : Globnames.global_reference -> (constr -> tactic) -> tactic +val pf_constr_of_global : Globnames.global_reference -> (Constr.constr -> tactic) -> tactic val elim_on_ba : (branch_assumptions -> tactic) -> branch_args -> tactic val case_on_ba : (branch_assumptions -> tactic) -> branch_args -> tactic @@ -230,13 +231,13 @@ module New : sig val nLastDecls : ([ `NF ], 'r) Proofview.Goal.t -> int -> Context.Named.t - val ifOnHyp : (identifier * EConstr.types -> bool) -> + val ifOnHyp : (identifier * types -> bool) -> (identifier -> unit Proofview.tactic) -> (identifier -> unit Proofview.tactic) -> identifier -> unit Proofview.tactic val onNthHypId : int -> (identifier -> unit tactic) -> unit tactic val onLastHypId : (identifier -> unit tactic) -> unit tactic - val onLastHyp : (EConstr.constr -> unit tactic) -> unit tactic + val onLastHyp : (constr -> unit tactic) -> unit tactic val onLastDecl : (Context.Named.Declaration.t -> unit tactic) -> unit tactic val onHyps : ([ `NF ], Context.Named.t) Proofview.Goal.enter -> @@ -253,18 +254,18 @@ module New : sig val elimination_then : (branch_args -> unit Proofview.tactic) -> - EConstr.constr -> unit Proofview.tactic + constr -> unit Proofview.tactic val case_then_using : or_and_intro_pattern option -> (branch_args -> unit Proofview.tactic) -> - EConstr.constr option -> pinductive -> EConstr.constr * EConstr.types -> unit Proofview.tactic + constr option -> pinductive -> constr * types -> unit Proofview.tactic val case_nodep_then_using : or_and_intro_pattern option -> (branch_args -> unit Proofview.tactic) -> - EConstr.constr option -> pinductive -> EConstr.constr * EConstr.types -> unit Proofview.tactic + constr option -> pinductive -> constr * types -> unit Proofview.tactic val elim_on_ba : (branch_assumptions -> unit Proofview.tactic) -> branch_args -> unit Proofview.tactic val case_on_ba : (branch_assumptions -> unit Proofview.tactic) -> branch_args -> unit Proofview.tactic - val pf_constr_of_global : Globnames.global_reference -> (constr -> unit Proofview.tactic) -> unit Proofview.tactic + val pf_constr_of_global : Globnames.global_reference -> (Constr.constr -> unit Proofview.tactic) -> unit Proofview.tactic end diff --git a/tactics/tactics.ml b/tactics/tactics.ml index f79f7f1a8..e4dd9eea2 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -2821,7 +2821,6 @@ let generalize_goal_gen env sigma ids i ((occs,c,b),na) t cl = let dummy_prod = it_mkProd_or_LetIn mkProp decls in let newdecls,_ = decompose_prod_n_assum sigma i (subst_term_gen sigma EConstr.eq_constr_nounivs c dummy_prod) in let cl',sigma' = subst_closed_term_occ env sigma (AtOccs occs) c (it_mkProd_or_LetIn cl newdecls) in - let cl' = EConstr.of_constr cl' in let na = generalized_name sigma c t ids cl' na in let decl = match b with | None -> local_assum (na,t) -- cgit v1.2.3 From 01849481fbabc3a3fa6c483e703996b01e37fca5 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sat, 26 Nov 2016 01:25:11 +0100 Subject: Removing compatibility layers from Tacticals --- plugins/cc/cctac.ml | 5 ++--- plugins/firstorder/instances.ml | 12 +++++++----- plugins/firstorder/rules.ml | 32 +++++++++++++++++++------------- plugins/fourier/fourierR.ml | 2 +- tactics/equality.ml | 2 -- tactics/tacticals.ml | 17 ++++++++++------- tactics/tacticals.mli | 8 ++++---- tactics/tactics.ml | 6 +++--- 8 files changed, 46 insertions(+), 38 deletions(-) diff --git a/plugins/cc/cctac.ml b/plugins/cc/cctac.ml index 2ab4dced4..53c450116 100644 --- a/plugins/cc/cctac.ml +++ b/plugins/cc/cctac.ml @@ -241,10 +241,10 @@ let build_projection intype (cstr:pconstructor) special default gls= let _M =mkMeta let app_global f args k = - Tacticals.pf_constr_of_global (Lazy.force f) (fun fc -> k (mkApp (EConstr.of_constr fc, args))) + Tacticals.pf_constr_of_global (Lazy.force f) (fun fc -> k (mkApp (fc, args))) let new_app_global f args k = - Tacticals.New.pf_constr_of_global (Lazy.force f) (fun fc -> k (mkApp (EConstr.of_constr fc, args))) + Tacticals.New.pf_constr_of_global (Lazy.force f) (fun fc -> k (mkApp (fc, args))) let new_refine c = Proofview.V82.tactic (refine c) let refine c = refine c @@ -491,7 +491,6 @@ let congruence_tac depth l = let mk_eq f c1 c2 k = Tacticals.New.pf_constr_of_global (Lazy.force f) (fun fc -> - let fc = EConstr.of_constr fc in Proofview.Goal.enter { enter = begin fun gl -> let open Tacmach.New in let evm, ty = pf_apply type_of gl c1 in diff --git a/plugins/firstorder/instances.ml b/plugins/firstorder/instances.ml index 2881b5333..ef8172de4 100644 --- a/plugins/firstorder/instances.ml +++ b/plugins/firstorder/instances.ml @@ -104,7 +104,7 @@ let mk_open_instance id idc gl m t= let evmap=Refiner.project gl in let var_id= if id==dummy_id then dummy_bvid else - let typ=pf_unsafe_type_of gl (EConstr.of_constr idc) in + let typ=pf_unsafe_type_of gl idc in (* since we know we will get a product, reduction is not too expensive *) let (nam,_,_)=destProd (EConstr.Unsafe.to_constr (whd_all env evmap typ)) in @@ -127,6 +127,7 @@ let mk_open_instance id idc gl m t= (* tactics *) let left_instance_tac (inst,id) continue seq= + let open EConstr in match inst with Phantom dom-> if lookup (id,None) seq then @@ -137,8 +138,8 @@ let left_instance_tac (inst,id) continue seq= [Proofview.V82.of_tactic introf; pf_constr_of_global id (fun idc -> (fun gls-> Proofview.V82.of_tactic (generalize - [EConstr.of_constr (mkApp(idc, - [|mkVar (Tacmach.pf_nth_hyp_id gls 1)|]))]) gls)); + [mkApp(idc, + [|mkVar (Tacmach.pf_nth_hyp_id gls 1)|])]) gls)); Proofview.V82.of_tactic introf; tclSOLVE [wrap 1 false continue (deepen (record (id,None) seq))]]; @@ -152,18 +153,19 @@ let left_instance_tac (inst,id) continue seq= pf_constr_of_global id (fun idc -> fun gl-> let evmap,rc,ot = mk_open_instance id idc gl m t in + let ot = EConstr.of_constr ot in let gt= it_mkLambda_or_LetIn (mkApp(idc,[|ot|])) rc in - let gt = EConstr.of_constr gt in let evmap, _ = try Typing.type_of (pf_env gl) evmap gt with e when CErrors.noncritical e -> error "Untypable instance, maybe higher-order non-prenex quantification" in tclTHEN (Refiner.tclEVARS evmap) (Proofview.V82.of_tactic (generalize [gt])) gl) else + let t = EConstr.of_constr t in pf_constr_of_global id (fun idc -> - Proofview.V82.of_tactic (generalize [EConstr.of_constr (mkApp(idc,[|t|]))])) + Proofview.V82.of_tactic (generalize [mkApp(idc,[|t|])])) in tclTHENLIST [special_generalize; diff --git a/plugins/firstorder/rules.ml b/plugins/firstorder/rules.ml index 38dae0b20..36bd91ab6 100644 --- a/plugins/firstorder/rules.ml +++ b/plugins/firstorder/rules.ml @@ -59,16 +59,17 @@ let clear_global=function (* connection rules *) let axiom_tac t seq= - try pf_constr_of_global (find_left t seq) (fun c -> Proofview.V82.of_tactic (exact_no_check (EConstr.of_constr c))) + try pf_constr_of_global (find_left t seq) (fun c -> Proofview.V82.of_tactic (exact_no_check c)) with Not_found->tclFAIL 0 (Pp.str "No axiom link") let ll_atom_tac a backtrack id continue seq= + let open EConstr in tclIFTHENELSE (try tclTHENLIST [pf_constr_of_global (find_left a seq) (fun left -> pf_constr_of_global id (fun id -> - Proofview.V82.of_tactic (generalize [EConstr.of_constr (mkApp(id, [|left|]))]))); + Proofview.V82.of_tactic (generalize [(mkApp(id, [|left|]))]))); clear_global id; Proofview.V82.of_tactic intro] with Not_found->tclFAIL 0 (Pp.str "No link")) @@ -95,7 +96,7 @@ let left_and_tac ind backtrack id continue seq gls= let n=(construct_nhyps ind gls).(0) in tclIFTHENELSE (tclTHENLIST - [Proofview.V82.of_tactic (Tacticals.New.pf_constr_of_global id (EConstr.of_constr %> simplest_elim)); + [Proofview.V82.of_tactic (Tacticals.New.pf_constr_of_global id simplest_elim); clear_global id; tclDO n (Proofview.V82.of_tactic intro)]) (wrap n false continue seq) @@ -109,12 +110,12 @@ let left_or_tac ind backtrack id continue seq gls= tclDO n (Proofview.V82.of_tactic intro); wrap n false continue seq] in tclIFTHENSVELSE - (Proofview.V82.of_tactic (Tacticals.New.pf_constr_of_global id (EConstr.of_constr %> simplest_elim))) + (Proofview.V82.of_tactic (Tacticals.New.pf_constr_of_global id simplest_elim)) (Array.map f v) backtrack gls let left_false_tac id= - Proofview.V82.of_tactic (Tacticals.New.pf_constr_of_global id (EConstr.of_constr %> simplest_elim)) + Proofview.V82.of_tactic (Tacticals.New.pf_constr_of_global id simplest_elim) (* left arrow connective rules *) @@ -133,7 +134,7 @@ let ll_ind_tac (ind,u as indu) largs backtrack id continue seq gl= let head=mkApp ((lift p idc),[|capply|]) in EConstr.of_constr (it_mkLambda_or_LetIn head rc) in let lp=Array.length rcs in - let newhyps idc =List.init lp (myterm idc) in + let newhyps idc =List.init lp (myterm (EConstr.Unsafe.to_constr idc)) in tclIFTHENELSE (tclTHENLIST [pf_constr_of_global id (fun idc -> Proofview.V82.of_tactic (generalize (newhyps idc))); @@ -142,17 +143,22 @@ let ll_ind_tac (ind,u as indu) largs backtrack id continue seq gl= (wrap lp false continue seq) backtrack gl let ll_arrow_tac a b c backtrack id continue seq= + let open EConstr in + let open Vars in + let a = EConstr.of_constr a in + let b = EConstr.of_constr b in + let c = EConstr.of_constr c in let cc=mkProd(Anonymous,a,(lift 1 b)) in - let d idc =EConstr.of_constr (mkLambda (Anonymous,b, - mkApp (idc, [|mkLambda (Anonymous,(lift 1 a),(mkRel 2))|]))) in + let d idc = mkLambda (Anonymous,b, + mkApp (idc, [|mkLambda (Anonymous,(lift 1 a),(mkRel 2))|])) in tclORELSE - (tclTHENS (Proofview.V82.of_tactic (cut (EConstr.of_constr c))) + (tclTHENS (Proofview.V82.of_tactic (cut c)) [tclTHENLIST [Proofview.V82.of_tactic introf; clear_global id; wrap 1 false continue seq]; - tclTHENS (Proofview.V82.of_tactic (cut (EConstr.of_constr cc))) - [pf_constr_of_global id (fun c -> Proofview.V82.of_tactic (exact_no_check (EConstr.of_constr c))); + tclTHENS (Proofview.V82.of_tactic (cut cc)) + [pf_constr_of_global id (fun c -> Proofview.V82.of_tactic (exact_no_check c)); tclTHENLIST [pf_constr_of_global id (fun idc -> Proofview.V82.of_tactic (generalize [d idc])); clear_global id; @@ -177,7 +183,7 @@ let forall_tac backtrack continue seq= let left_exists_tac ind backtrack id continue seq gls= let n=(construct_nhyps ind gls).(0) in tclIFTHENELSE - (Proofview.V82.of_tactic (Tacticals.New.pf_constr_of_global id (EConstr.of_constr %> simplest_elim))) + (Proofview.V82.of_tactic (Tacticals.New.pf_constr_of_global id simplest_elim)) (tclTHENLIST [clear_global id; tclDO n (Proofview.V82.of_tactic intro); (wrap (n-1) false continue seq)]) @@ -191,9 +197,9 @@ let ll_forall_tac prod backtrack id continue seq= [Proofview.V82.of_tactic intro; pf_constr_of_global id (fun idc -> (fun gls-> + let open EConstr in let id0=pf_nth_hyp_id gls 1 in let term=mkApp(idc,[|mkVar(id0)|]) in - let term = EConstr.of_constr term in tclTHEN (Proofview.V82.of_tactic (generalize [term])) (Proofview.V82.of_tactic (clear [id0])) gls)); clear_global id; Proofview.V82.of_tactic intro; diff --git a/plugins/fourier/fourierR.ml b/plugins/fourier/fourierR.ml index bffca6223..ec73fccb5 100644 --- a/plugins/fourier/fourierR.ml +++ b/plugins/fourier/fourierR.ml @@ -616,7 +616,7 @@ let rec fourier () = (* TODO : Ring.polynom []*) (Proofview.tclUNIT ()) (Proofview.tclUNIT ()); Tacticals.New.pf_constr_of_global (cget coq_sym_eqT) (fun symeq -> - (Tacticals.New.tclTHEN (apply (EConstr.of_constr symeq)) + (Tacticals.New.tclTHEN (apply symeq) (apply (get coq_Rinv_1))))] ) diff --git a/tactics/equality.ml b/tactics/equality.ml index 7dcfd419e..c80cf4416 100644 --- a/tactics/equality.ml +++ b/tactics/equality.ml @@ -647,9 +647,7 @@ let replace_using_leibniz clause c1 c2 l2r unsafe try_prove_eq_opt = let sym = build_coq_eq_sym () in Tacticals.New.pf_constr_of_global sym (fun sym -> Tacticals.New.pf_constr_of_global e (fun e -> - let e = EConstr.of_constr e in let eq = applist (e, [t1;c1;c2]) in - let sym = EConstr.of_constr sym in tclTHENLAST (replace_core clause l2r eq) (tclFIRST diff --git a/tactics/tacticals.ml b/tactics/tacticals.ml index 89acc149c..9cf3c4187 100644 --- a/tactics/tacticals.ml +++ b/tactics/tacticals.ml @@ -11,6 +11,7 @@ open CErrors open Util open Names open Term +open EConstr open Termops open Declarations open Tacmach @@ -73,7 +74,7 @@ let nthDecl m gl = with Failure _ -> error "No such assumption." let nthHypId m gl = nthDecl m gl |> NamedDecl.get_id -let nthHyp m gl = EConstr.mkVar (nthHypId m gl) +let nthHyp m gl = mkVar (nthHypId m gl) let lastDecl gl = nthDecl 1 gl let lastHypId gl = nthHypId 1 gl @@ -83,7 +84,7 @@ let nLastDecls n gl = try List.firstn n (pf_hyps gl) with Failure _ -> error "Not enough hypotheses in the goal." -let nLastHypsId n gl = List.map NamedDecl.get_id (nLastDecls n gl) +let nLastHypsId n gl = List.map (NamedDecl.get_id) (nLastDecls n gl) let nLastHyps n gl = List.map mkVar (nLastHypsId n gl) let onNthDecl m tac gl = tac (nthDecl m gl) gl @@ -147,9 +148,9 @@ let ifOnHyp pred tac1 tac2 id gl = type branch_args = { ity : pinductive; (* the type we were eliminating on *) - largs : EConstr.constr list; (* its arguments *) + largs : constr list; (* its arguments *) branchnum : int; (* the branch number *) - pred : EConstr.constr; (* the predicate we used *) + pred : constr; (* the predicate we used *) nassums : int; (* number of assumptions/letin to be introduced *) branchsign : bool list; (* the signature of the branch. true=assumption, false=let-in *) @@ -226,6 +227,7 @@ let compute_induction_names = compute_induction_names_gen true (* Compute the let-in signature of case analysis or standard induction scheme *) let compute_constructor_signatures isrec ((_,k as ity),u) = + let open Term in let rec analrec c recargs = match kind_of_term c, recargs with | Prod (_,_,c), recarg::rest -> @@ -263,7 +265,7 @@ let pf_with_evars glsev k gls = tclTHEN (Refiner.tclEVARS evd) (k a) gls let pf_constr_of_global gr k = - pf_with_evars (fun gls -> pf_apply Evd.fresh_global gls gr) k + pf_with_evars (fun gls -> on_snd EConstr.of_constr (pf_apply Evd.fresh_global gls gr)) k (* computing the case/elim combinators *) @@ -565,7 +567,7 @@ module New = struct let gl = Proofview.Goal.assume gl in nthDecl m gl |> NamedDecl.get_id let nthHyp m gl = - EConstr.mkVar (nthHypId m gl) + mkVar (nthHypId m gl) let onNthHypId m tac = Proofview.Goal.enter { enter = begin fun gl -> tac (nthHypId m gl) end } @@ -657,7 +659,7 @@ module New = struct let elimclause' = match predicate with | None -> elimclause' - | Some p -> clenv_unify ~flags Reduction.CONV (EConstr.mkMeta pmv) p elimclause' + | Some p -> clenv_unify ~flags Reduction.CONV (mkMeta pmv) p elimclause' in let clenv' = Tacmach.New.of_old (clenv_unique_resolver ~flags elimclause') gl in let after_tac i = @@ -726,6 +728,7 @@ module New = struct let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in let (sigma, c) = Evd.fresh_global env sigma ref in + let c = EConstr.of_constr c in Proofview.Unsafe.tclEVARS sigma <*> (tac c) end } diff --git a/tactics/tacticals.mli b/tactics/tacticals.mli index ba5452e33..2b07d937e 100644 --- a/tactics/tacticals.mli +++ b/tactics/tacticals.mli @@ -65,14 +65,14 @@ val onLastHypId : (Id.t -> tactic) -> tactic val onLastHyp : (constr -> tactic) -> tactic val onLastDecl : (Context.Named.Declaration.t -> tactic) -> tactic val onNLastHypsId : int -> (Id.t list -> tactic) -> tactic -val onNLastHyps : int -> (Constr.constr list -> tactic) -> tactic +val onNLastHyps : int -> (constr list -> tactic) -> tactic val onNLastDecls : int -> (Context.Named.t -> tactic) -> tactic val lastHypId : goal sigma -> Id.t val lastHyp : goal sigma -> constr val lastDecl : goal sigma -> Context.Named.Declaration.t val nLastHypsId : int -> goal sigma -> Id.t list -val nLastHyps : int -> goal sigma -> Constr.constr list +val nLastHyps : int -> goal sigma -> constr list val nLastDecls : int -> goal sigma -> Context.Named.t val afterHyp : Id.t -> goal sigma -> Context.Named.t @@ -135,7 +135,7 @@ val elimination_sort_of_hyp : Id.t -> goal sigma -> sorts_family val elimination_sort_of_clause : Id.t option -> goal sigma -> sorts_family val pf_with_evars : (goal sigma -> Evd.evar_map * 'a) -> ('a -> tactic) -> tactic -val pf_constr_of_global : Globnames.global_reference -> (Constr.constr -> tactic) -> tactic +val pf_constr_of_global : Globnames.global_reference -> (constr -> tactic) -> tactic val elim_on_ba : (branch_assumptions -> tactic) -> branch_args -> tactic val case_on_ba : (branch_assumptions -> tactic) -> branch_args -> tactic @@ -267,5 +267,5 @@ module New : sig val elim_on_ba : (branch_assumptions -> unit Proofview.tactic) -> branch_args -> unit Proofview.tactic val case_on_ba : (branch_assumptions -> unit Proofview.tactic) -> branch_args -> unit Proofview.tactic - val pf_constr_of_global : Globnames.global_reference -> (Constr.constr -> unit Proofview.tactic) -> unit Proofview.tactic + val pf_constr_of_global : Globnames.global_reference -> (constr -> unit Proofview.tactic) -> unit Proofview.tactic end diff --git a/tactics/tactics.ml b/tactics/tactics.ml index e4dd9eea2..10582288c 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -4742,7 +4742,7 @@ let symmetry_red allowred = | Some eq_data,_,_ -> Tacticals.New.tclTHEN (convert_concl_no_check concl DEFAULTcast) - (Tacticals.New.pf_constr_of_global eq_data.sym (EConstr.of_constr %> apply)) + (Tacticals.New.pf_constr_of_global eq_data.sym apply) | None,eq,eq_kind -> prove_symmetry eq eq_kind end } @@ -4838,8 +4838,8 @@ let transitivity_red allowred t = Tacticals.New.tclTHEN (convert_concl_no_check concl DEFAULTcast) (match t with - | None -> Tacticals.New.pf_constr_of_global eq_data.trans (EConstr.of_constr %> eapply) - | Some t -> Tacticals.New.pf_constr_of_global eq_data.trans (fun trans -> apply_list [EConstr.of_constr trans;t])) + | None -> Tacticals.New.pf_constr_of_global eq_data.trans eapply + | Some t -> Tacticals.New.pf_constr_of_global eq_data.trans (fun trans -> apply_list [trans;t])) | None,eq,eq_kind -> match t with | None -> Tacticals.New.tclZEROMSG (str"etransitivity not supported for this relation.") -- cgit v1.2.3 From c8c8ccdaaffefdbd3d78c844552a08bcb7b4f915 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sat, 26 Nov 2016 02:12:40 +0100 Subject: Evar-normalizing functions now act on EConstrs. --- engine/evarutil.ml | 31 +++++------ engine/evarutil.mli | 4 +- engine/proofview.ml | 22 ++++---- engine/proofview.mli | 8 +-- engine/termops.mli | 4 +- ide/ide_slave.ml | 5 +- ltac/rewrite.ml | 23 ++++----- ltac/tacinterp.ml | 1 - plugins/decl_mode/decl_proof_instr.ml | 1 + plugins/derive/derive.ml | 5 +- plugins/funind/functional_principles_proofs.ml | 2 +- plugins/funind/functional_principles_types.ml | 2 +- plugins/funind/invfun.ml | 4 +- plugins/funind/recdef.ml | 6 +-- plugins/ssrmatching/ssrmatching.ml4 | 25 +++++---- pretyping/cases.ml | 8 +-- pretyping/coercion.ml | 1 - pretyping/evarconv.ml | 6 +-- pretyping/evardefine.ml | 5 +- pretyping/evarsolve.ml | 4 +- pretyping/pretype_errors.ml | 37 +++++++------- pretyping/pretype_errors.mli | 4 +- pretyping/pretyping.ml | 3 -- pretyping/reductionops.mli | 4 +- pretyping/typing.ml | 4 +- pretyping/unification.ml | 12 ++--- printing/printer.ml | 8 +-- proofs/pfedit.ml | 4 +- proofs/pfedit.mli | 10 ++-- proofs/proof.mli | 6 +-- proofs/proof_global.ml | 5 +- proofs/proof_global.mli | 4 +- proofs/tacmach.ml | 4 +- stm/lemmas.ml | 6 +-- stm/lemmas.mli | 6 +-- stm/stm.ml | 4 +- tactics/auto.ml | 2 +- tactics/class_tactics.ml | 6 +-- tactics/contradiction.ml | 3 +- tactics/eauto.ml | 3 +- tactics/equality.ml | 2 +- tactics/hints.ml | 3 +- tactics/leminv.ml | 2 - tactics/tactics.ml | 4 +- toplevel/auto_ind_decl.ml | 4 +- toplevel/classes.ml | 4 +- toplevel/command.ml | 71 +++++++++++++++----------- toplevel/himsg.ml | 56 ++++++++++---------- toplevel/obligations.ml | 4 +- toplevel/vernacentries.ml | 8 +-- 50 files changed, 233 insertions(+), 227 deletions(-) diff --git a/engine/evarutil.ml b/engine/evarutil.ml index 8940be09d..42620c0b5 100644 --- a/engine/evarutil.ml +++ b/engine/evarutil.ml @@ -99,16 +99,21 @@ let rec whd_evar sigma c = if u' == u then c else mkConstructU (co, u') | _ -> c -let rec nf_evar sigma t = Constr.map (fun t -> nf_evar sigma t) (whd_evar sigma t) -let e_nf_evar sigma t = EConstr.of_constr (nf_evar sigma (EConstr.Unsafe.to_constr t)) +(** Term exploration up to instantiation. *) +let kind_of_term_upto sigma t = + Constr.kind (whd_evar sigma t) + +let rec nf_evar0 sigma t = Constr.map (fun t -> nf_evar0 sigma t) (whd_evar sigma t) +let whd_evar sigma c = EConstr.of_constr (whd_evar sigma (EConstr.Unsafe.to_constr c)) +let nf_evar sigma c = EConstr.of_constr (nf_evar0 sigma (EConstr.Unsafe.to_constr c)) let j_nf_evar sigma j = - { uj_val = e_nf_evar sigma j.uj_val; - uj_type = e_nf_evar sigma j.uj_type } + { uj_val = nf_evar sigma j.uj_val; + uj_type = nf_evar sigma j.uj_type } let jl_nf_evar sigma jl = List.map (j_nf_evar sigma) jl let jv_nf_evar sigma = Array.map (j_nf_evar sigma) let tj_nf_evar sigma {utj_val=v;utj_type=t} = - {utj_val=e_nf_evar sigma v;utj_type=t} + {utj_val=nf_evar sigma v;utj_type=t} let nf_evars_universes evm = Universes.nf_evars_and_universes_opt_subst (safe_evar_value evm) @@ -125,23 +130,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 + Context.Rel.map (nf_evar0 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 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 @@ -527,7 +532,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 @@ -582,7 +587,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 @@ -751,10 +756,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 diff --git a/engine/evarutil.mli b/engine/evarutil.mli index ea9599c8b..67de18abc 100644 --- a/engine/evarutil.mli +++ b/engine/evarutil.mli @@ -139,8 +139,8 @@ val judge_of_new_Type : 'r Sigma.t -> (unsafe_judgment, 'r) Sigma.sigma (** [flush_and_check_evars] raise [Uninstantiated_evar] if an evar remains uninstantiated; [nf_evar] leaves uninstantiated evars as is *) -val whd_evar : evar_map -> constr -> constr -val nf_evar : evar_map -> constr -> constr +val whd_evar : evar_map -> EConstr.constr -> EConstr.constr +val nf_evar : evar_map -> EConstr.constr -> EConstr.constr val j_nf_evar : evar_map -> EConstr.unsafe_judgment -> EConstr.unsafe_judgment val jl_nf_evar : evar_map -> EConstr.unsafe_judgment list -> EConstr.unsafe_judgment list diff --git a/engine/proofview.ml b/engine/proofview.ml index 9e5e9c7da..ab72cc405 100644 --- a/engine/proofview.ml +++ b/engine/proofview.ml @@ -24,7 +24,7 @@ type proofview = Proofview_monad.proofview (* 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 = (Term.constr * Term.types) list +type entry = (EConstr.constr * EConstr.types) list (** Returns a stylised view of a proofview for use by, for instance, ide-s. *) @@ -37,15 +37,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)); @@ -56,7 +57,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 () @@ -71,11 +72,10 @@ let dependent_init = | 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 (EConstr.of_constr typ) in - let econstr = EConstr.Unsafe.to_constr econstr in + let Sigma (econstr, sigma, _) = Evarutil.new_evar env sigma ~src ~store typ in + let (gl, _) = EConstr.destEvar (Sigma.to_evar_map sigma) econstr in let sigma = Sigma.to_evar_map sigma 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 @@ -1222,12 +1222,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) diff --git a/engine/proofview.mli b/engine/proofview.mli index 2350592a2..9f10e874a 100644 --- a/engine/proofview.mli +++ b/engine/proofview.mli @@ -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 * EConstr.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 * EConstr.types * (Evd.evar_map -> EConstr.constr -> telescope) (** Like {!init}, but goals are allowed to be dependent on one another. Dependencies between goals is represented with the type @@ -69,8 +69,8 @@ val finished : proofview -> bool (** Returns the current [evar] state. *) val return : proofview -> Evd.evar_map -val partial_proof : entry -> proofview -> constr list -val initial_goals : entry -> (constr * types) list +val partial_proof : entry -> proofview -> EConstr.constr list +val initial_goals : entry -> (EConstr.constr * EConstr.types) list diff --git a/engine/termops.mli b/engine/termops.mli index 3f924cfa1..196962354 100644 --- a/engine/termops.mli +++ b/engine/termops.mli @@ -12,8 +12,8 @@ open Pp open Names open Term -open EConstr open Environ +open EConstr (** printers *) val print_sort : sorts -> std_ppcmds @@ -283,4 +283,4 @@ val on_judgment_type : ('t -> 't) -> ('c, 't) punsafe_judgment -> ('c, 't) puns (** {6 Functions to deal with impossible cases } *) val set_impossible_default_clause : (unit -> (Constr.constr * Constr.types) Univ.in_universe_context_set) -> unit -val coq_unit_judge : unit -> EConstr.unsafe_judgment Univ.in_universe_context_set +val coq_unit_judge : unit -> unsafe_judgment Univ.in_universe_context_set diff --git a/ide/ide_slave.ml b/ide/ide_slave.ml index fcba01353..8fc5547ec 100644 --- a/ide/ide_slave.ml +++ b/ide/ide_slave.ml @@ -191,12 +191,11 @@ let process_goal sigma g = let min_env = Environ.reset_context env in let id = Goal.uid g in let ccl = - let norm_constr = Reductionops.nf_evar sigma (EConstr.Unsafe.to_constr (Goal.V82.concl sigma g)) in - let norm_constr = EConstr.of_constr norm_constr in + let norm_constr = Reductionops.nf_evar sigma (Goal.V82.concl sigma g) in Richpp.richpp_of_pp (pr_goal_concl_style_env env sigma norm_constr) in let process_hyp d (env,l) = - let d = CompactedDecl.map_constr (Reductionops.nf_evar sigma) d in + let d = CompactedDecl.map_constr (fun c -> EConstr.Unsafe.to_constr (Reductionops.nf_evar sigma (EConstr.of_constr c))) d in let d' = CompactedDecl.to_named_context d in (List.fold_right Environ.push_named d' env, (Richpp.richpp_of_pp (pr_compacted_decl env sigma d)) :: l) in diff --git a/ltac/rewrite.ml b/ltac/rewrite.ml index 715929c56..c942e8e92 100644 --- a/ltac/rewrite.ml +++ b/ltac/rewrite.ml @@ -57,9 +57,6 @@ let nlocal_def (na, b, t) = let inj = EConstr.Unsafe.to_constr in NamedDecl.LocalDef (na, inj b, inj t) -let nf_evar sigma c = - EConstr.of_constr (Evarutil.nf_evar sigma (EConstr.Unsafe.to_constr c)) - (** Typeclass-based generalized rewriting. *) (** Constants used by the tactic. *) @@ -663,7 +660,7 @@ let solve_remaining_by env sigma holes by = (** Evar should not be defined, but just in case *) | Some evi -> let env = Environ.reset_with_named_context evi.evar_hyps env in - let ty = evi.evar_concl in + let ty = EConstr.of_constr evi.evar_concl in let c, sigma = Pfedit.refine_by_tactic env sigma ty solve_tac in Evd.define evk c sigma in @@ -738,7 +735,7 @@ let unify_eqn (car, rel, prf, c1, c2, holes, sort) l2r flags env (sigma, cstrs) let sigma = Typeclasses.resolve_typeclasses ~filter:(no_constraints cstrs) ~fail:true env sigma in let evd = solve_remaining_by env sigma holes by in - let nf c = nf_evar evd (Reductionops.nf_meta evd c) in + let nf c = Reductionops.nf_evar evd (Reductionops.nf_meta evd c) in let c1 = nf c1 and c2 = nf c2 and rew_car = nf car and rel = nf rel and prf = nf prf in @@ -1449,7 +1446,7 @@ module Strategies = in try let sigma = Unification.w_unify env sigma CONV ~flags:(Unification.elim_flags ()) unfolded t in - let c' = nf_evar sigma c in + let c' = Reductionops.nf_evar sigma c in state, Success { rew_car = ty; rew_from = t; rew_to = c'; rew_prf = RewCast DEFAULTcast; rew_evars = (sigma, snd evars) } @@ -1522,7 +1519,7 @@ let cl_rewrite_clause_aux ?(abs=None) strat env avoid sigma concl is_hyp : resul | Success res -> let (_, cstrs) = res.rew_evars in let evars' = solve_constraints env res.rew_evars in - let newt = nf_evar evars' res.rew_to in + let newt = Reductionops.nf_evar evars' res.rew_to in let evars = (* Keep only original evars (potentially instantiated) and goal evars, the rest has been defined and substituted already. *) Evar.Set.fold @@ -1537,13 +1534,13 @@ let cl_rewrite_clause_aux ?(abs=None) strat env avoid sigma concl is_hyp : resul let res = match res.rew_prf with | RewCast c -> None | RewPrf (rel, p) -> - let p = nf_zeta env evars' (nf_evar evars' p) in + let p = nf_zeta env evars' (Reductionops.nf_evar evars' p) in let term = match abs with | None -> p | Some (t, ty) -> - let t = nf_evar evars' t in - let ty = nf_evar evars' ty in + let t = Reductionops.nf_evar evars' t in + let ty = Reductionops.nf_evar evars' ty in mkApp (mkLambda (Name (Id.of_string "lemma"), ty, p), [| t |]) in let proof = match is_hyp with @@ -2024,7 +2021,7 @@ let add_morphism_infer glob m n = let hook = Lemmas.mk_hook hook in Flags.silently (fun () -> - Lemmas.start_proof instance_id kind (Evd.from_ctx uctx) instance hook; + Lemmas.start_proof instance_id kind (Evd.from_ctx uctx) (EConstr.of_constr instance) hook; ignore (Pfedit.by (Tacinterp.interp tac))) () let add_morphism glob binders m s n = @@ -2085,7 +2082,7 @@ let unification_rewrite l2r c1 c2 sigma prf car rel but env = ~flags:rewrite_conv_unif_flags env sigma ((if l2r then c1 else c2),but) in - let nf c = nf_evar sigma c in + let nf c = Reductionops.nf_evar sigma c in let c1 = if l2r then nf c' else nf c1 and c2 = if l2r then nf c2 else nf c' and car = nf car and rel = nf rel in @@ -2104,7 +2101,7 @@ let get_hyp gl (c,l) clause l2r = let sigma, hi = decompose_applied_relation env evars (c,l) in let but = match clause with | Some id -> Tacmach.New.pf_get_hyp_typ id gl - | None -> nf_evar evars (Tacmach.New.pf_concl gl) + | None -> Reductionops.nf_evar evars (Tacmach.New.pf_concl gl) in unification_rewrite l2r hi.c1 hi.c2 sigma hi.prf hi.car hi.rel but env diff --git a/ltac/tacinterp.ml b/ltac/tacinterp.ml index 36a0336bc..62d0cc529 100644 --- a/ltac/tacinterp.ml +++ b/ltac/tacinterp.ml @@ -2101,7 +2101,6 @@ let interp_redexp env sigma r = let _ = let eval ty env sigma lfun arg = - let ty = EConstr.Unsafe.to_constr ty in let ist = { lfun = lfun; extra = TacStore.empty; } in if Genarg.has_type arg (glbwit wit_tactic) then let tac = Genarg.out_gen (glbwit wit_tactic) arg in diff --git a/plugins/decl_mode/decl_proof_instr.ml b/plugins/decl_mode/decl_proof_instr.ml index 6a0ec3968..da971fffb 100644 --- a/plugins/decl_mode/decl_proof_instr.ml +++ b/plugins/decl_mode/decl_proof_instr.ml @@ -1506,6 +1506,7 @@ let rec postprocess pts instr = | Pend (B_elim ET_Induction) -> begin let pfterm = List.hd (Proof.partial_proof pts) in + let pfterm = EConstr.Unsafe.to_constr pfterm in let { it = gls ; sigma = sigma } = Proof.V82.subgoals pts in let env = try Goal.V82.env sigma (List.hd gls) diff --git a/plugins/derive/derive.ml b/plugins/derive/derive.ml index f23f4ce7d..12d7f0660 100644 --- a/plugins/derive/derive.ml +++ b/plugins/derive/derive.ml @@ -28,16 +28,17 @@ let start_deriving f suchthat lemma = (* spiwack: I don't know what the rigidity flag does, picked the one that looked the most general. *) let (sigma,f_type_sort) = Evd.new_sort_variable Evd.univ_flexible_alg sigma in - let f_type_type = Term.mkSort f_type_sort in + let f_type_type = EConstr.mkSort f_type_sort in (** create the initial goals for the proof: |- Type ; |- ?1 ; f:=?2 |- suchthat *) let goals = let open Proofview in TCons ( env , sigma , f_type_type , (fun sigma f_type -> TCons ( env , sigma , f_type , (fun sigma ef -> + let f_type = EConstr.Unsafe.to_constr f_type in + let ef = EConstr.Unsafe.to_constr ef in let env' = Environ.push_named (LocalDef (f, ef, f_type)) env in let evdref = ref sigma in let suchthat = Constrintern.interp_type_evars env' evdref suchthat in - let suchthat = EConstr.Unsafe.to_constr suchthat in TCons ( env' , !evdref , suchthat , (fun sigma _ -> TNil sigma)))))) in diff --git a/plugins/funind/functional_principles_proofs.ml b/plugins/funind/functional_principles_proofs.ml index f4fa61a22..91b17b9a4 100644 --- a/plugins/funind/functional_principles_proofs.ml +++ b/plugins/funind/functional_principles_proofs.ml @@ -1007,7 +1007,7 @@ let generate_equation_lemma evd fnames f fun_num nb_params nb_args rec_args_num (mk_equation_id f_id) (Decl_kinds.Global, Flags.is_universe_polymorphism (), (Decl_kinds.Proof Decl_kinds.Theorem)) evd - (EConstr.Unsafe.to_constr lemma_type) + lemma_type (Lemmas.mk_hook (fun _ _ -> ())); ignore (Pfedit.by (Proofview.V82.tactic prove_replacement)); Lemmas.save_proof (Vernacexpr.(Proved(Transparent,None))); diff --git a/plugins/funind/functional_principles_types.ml b/plugins/funind/functional_principles_types.ml index ba01b3b04..d0d44b34b 100644 --- a/plugins/funind/functional_principles_types.ml +++ b/plugins/funind/functional_principles_types.ml @@ -294,7 +294,7 @@ let build_functional_principle (evd:Evd.evar_map ref) interactive_proof old_prin new_princ_name (Decl_kinds.Global,Flags.is_universe_polymorphism (),(Decl_kinds.Proof Decl_kinds.Theorem)) !evd - new_principle_type + (EConstr.of_constr new_principle_type) hook ; (* let _tim1 = System.get_time () in *) diff --git a/plugins/funind/invfun.ml b/plugins/funind/invfun.ml index 5cbec7743..dcec2cb74 100644 --- a/plugins/funind/invfun.ml +++ b/plugins/funind/invfun.ml @@ -846,7 +846,7 @@ let derive_correctness make_scheme functional_induction (funs: pconstant list) ( lem_id (Decl_kinds.Global,Flags.is_universe_polymorphism (),((Decl_kinds.Proof Decl_kinds.Theorem))) !evd - (EConstr.Unsafe.to_constr typ) + typ (Lemmas.mk_hook (fun _ _ -> ())); ignore (Pfedit.by (Proofview.V82.tactic (observe_tac ("prove correctness ("^(Id.to_string f_id)^")") @@ -908,7 +908,7 @@ let derive_correctness make_scheme functional_induction (funs: pconstant list) ( let lem_id = mk_complete_id f_id in Lemmas.start_proof lem_id (Decl_kinds.Global,Flags.is_universe_polymorphism (),(Decl_kinds.Proof Decl_kinds.Theorem)) sigma - (EConstr.Unsafe.to_constr (fst lemmas_types_infos.(i))) + (fst lemmas_types_infos.(i)) (Lemmas.mk_hook (fun _ _ -> ())); ignore (Pfedit.by (Proofview.V82.tactic (observe_tac ("prove completeness ("^(Id.to_string f_id)^")") diff --git a/plugins/funind/recdef.ml b/plugins/funind/recdef.ml index adbdb1eb7..56c6ab054 100644 --- a/plugins/funind/recdef.ml +++ b/plugins/funind/recdef.ml @@ -1374,7 +1374,7 @@ let open_new_goal build_proof sigma using_lemmas ref_ goal_name (gls_type,decomp Lemmas.start_proof na (Decl_kinds.Global, false (* FIXME *), Decl_kinds.Proof Decl_kinds.Lemma) - sigma (EConstr.Unsafe.to_constr gls_type) + sigma gls_type (Lemmas.mk_hook hook); if Indfun_common.is_strict_tcc () then @@ -1421,7 +1421,7 @@ let com_terminate let (evmap, env) = Lemmas.get_current_context() in Lemmas.start_proof thm_name (Global, false (* FIXME *), Proof Lemma) ~sign:(Environ.named_context_val env) - ctx (compute_terminate_type nb_args fonctional_ref) hook; + ctx (EConstr.of_constr (compute_terminate_type nb_args fonctional_ref)) hook; ignore (by (Proofview.V82.tactic (observe_tac (str "starting_tac") tac_start))); ignore (by (Proofview.V82.tactic (observe_tac (str "whole_start") (whole_start tac_end nb_args is_mes fonctional_ref @@ -1476,7 +1476,7 @@ let (com_eqn : int -> Id.t -> (Lemmas.start_proof eq_name (Global, false, Proof Lemma) ~sign:(Environ.named_context_val env) evmap - equation_lemma_type + (EConstr.of_constr equation_lemma_type) (Lemmas.mk_hook (fun _ _ -> ())); ignore (by (Proofview.V82.tactic (start_equation f_ref terminate_ref diff --git a/plugins/ssrmatching/ssrmatching.ml4 b/plugins/ssrmatching/ssrmatching.ml4 index eb5f401e3..bf3e2ac1c 100644 --- a/plugins/ssrmatching/ssrmatching.ml4 +++ b/plugins/ssrmatching/ssrmatching.ml4 @@ -178,6 +178,9 @@ let mk_lterm = mk_term ' ' let pf_type_of gl t = let sigma, ty = pf_type_of gl t in re_sig (sig_it gl) sigma, ty +let nf_evar sigma c = + EConstr.Unsafe.to_constr (Evarutil.nf_evar sigma (EConstr.of_constr c)) + (* }}} *) (** Profiling {{{ *************************************************************) @@ -780,13 +783,13 @@ let on_instance, instances = let rec uniquize = function | [] -> [] | (sigma,_,{ up_f = f; up_a = a; up_t = t } as x) :: xs -> - let t = Reductionops.nf_evar sigma t in - let f = Reductionops.nf_evar sigma f in - let a = Array.map (Reductionops.nf_evar sigma) a in + let t = nf_evar sigma t in + let f = nf_evar sigma f in + let a = Array.map (nf_evar sigma) a in let neq (sigma1,_,{ up_f = f1; up_a = a1; up_t = t1 }) = - let t1 = Reductionops.nf_evar sigma1 t1 in - let f1 = Reductionops.nf_evar sigma1 f1 in - let a1 = Array.map (Reductionops.nf_evar sigma1) a1 in + let t1 = nf_evar sigma1 t1 in + let f1 = nf_evar sigma1 f1 in + let a1 = Array.map (nf_evar sigma1) a1 in not (Term.eq_constr t t1 && Term.eq_constr f f1 && CArray.for_all2 Term.eq_constr a a1) in x :: uniquize (List.filter neq xs) in @@ -1138,7 +1141,7 @@ let interp_pattern ?wit_ssrpatternarg ist gl red redty = if k = h_k || List.mem k acc || Evd.mem sigma0 k then acc else (update k; k::acc) | _ -> fold_constr aux acc t in - aux [] (Evarutil.nf_evar sigma rp) in + aux [] (nf_evar sigma rp) in let sigma = List.fold_left (fun sigma e -> if Evd.is_defined sigma e then sigma else (* clear may be recursive *) @@ -1195,7 +1198,7 @@ let interp_pattern ?wit_ssrpatternarg ist gl red redty = let sigma, rp = interp_term ist gl rp in let _, h, _, rp = destLetIn rp in let sigma = cleanup_XinE h x rp sigma in - let rp = subst1 h (Evarutil.nf_evar sigma rp) in + let rp = subst1 h (nf_evar sigma rp) in sigma, mk h rp | E_In_X_In_T(e, x, rp) | E_As_X_In_T (e, x, rp) -> let mk e x p = @@ -1204,7 +1207,7 @@ let interp_pattern ?wit_ssrpatternarg ist gl red redty = let sigma, rp = interp_term ist gl rp in let _, h, _, rp = destLetIn rp in let sigma = cleanup_XinE h x rp sigma in - let rp = subst1 h (Evarutil.nf_evar sigma rp) in + let rp = subst1 h (nf_evar sigma rp) in let sigma, e = interp_term ist (re_sig (sig_it gl) sigma) e in sigma, mk e h rp ;; @@ -1220,7 +1223,7 @@ let noindex = Some(false,[]) (* calls do_subst on every sub-term identified by (pattern,occ) *) let eval_pattern ?raise_NoMatch env0 sigma0 concl0 pattern occ do_subst = - let fs sigma x = Reductionops.nf_evar sigma x in + let fs sigma x = nf_evar sigma x in let pop_evar sigma e p = let { Evd.evar_body = e_body } as e_def = Evd.find sigma e in let e_body = match e_body with Evar_defined c -> c @@ -1307,7 +1310,7 @@ let redex_of_pattern ?(resolve_typeclasses=false) env (sigma, p) = let sigma = if not resolve_typeclasses then sigma else Typeclasses.resolve_typeclasses ~fail:false env sigma in - Reductionops.nf_evar sigma e, Evd.evar_universe_context sigma + nf_evar sigma e, Evd.evar_universe_context sigma let fill_occ_pattern ?raise_NoMatch env sigma cl pat occ h = let do_make_rel, occ = diff --git a/pretyping/cases.ml b/pretyping/cases.ml index eea94f021..0e4c6619b 100644 --- a/pretyping/cases.ml +++ b/pretyping/cases.ml @@ -362,7 +362,7 @@ let coerce_row typing_fun evdref env pats (tomatch,(_,indopt)) = let j = typing_fun tycon env evdref tomatch in let evd, j = Coercion.inh_coerce_to_base (loc_of_glob_constr tomatch) env !evdref j in evdref := evd; - let typ = EConstr.of_constr (nf_evar !evdref (EConstr.Unsafe.to_constr j.uj_type)) in + let typ = nf_evar !evdref j.uj_type in let t = try try_find_ind env !evdref typ realnames with Not_found -> @@ -1145,7 +1145,7 @@ let postprocess_dependencies evd tocheck brs tomatch pred deps cs = let rec aux k brs tomatch pred tocheck deps = match deps, tomatch with | [], _ -> brs,tomatch,pred,[] | n::deps, Abstract (i,d) :: tomatch -> - let d = map_constr (nf_evar evd) d in + let d = map_constr (fun c -> EConstr.Unsafe.to_constr (nf_evar evd (EConstr.of_constr c))) d in let is_d = match d with LocalAssum _ -> false | LocalDef _ -> true in if is_d || List.exists (fun c -> dependent_decl evd (lift k c) d) tocheck && Array.exists (is_dependent_branch evd k) brs then @@ -2008,9 +2008,9 @@ let prepare_predicate loc typing_fun env sigma tomatchs arsign tycon pred = let envar = List.fold_right push_rel_context arsign env in let sigma, newt = new_sort_variable univ_flexible_alg sigma in let evdref = ref sigma in - let predcclj = typing_fun (mk_tycon (EConstr.mkSort newt)) envar evdref rtntyp in + let predcclj = typing_fun (mk_tycon (mkSort newt)) envar evdref rtntyp in let sigma = !evdref in - let predccl = EConstr.of_constr (nf_evar sigma (EConstr.Unsafe.to_constr predcclj.uj_val)) in + let predccl = nf_evar sigma predcclj.uj_val in [sigma, predccl] in List.map diff --git a/pretyping/coercion.ml b/pretyping/coercion.ml index ead44cee2..91f53a886 100644 --- a/pretyping/coercion.ml +++ b/pretyping/coercion.ml @@ -417,7 +417,6 @@ let inh_tosort_force loc env evd j = try let t,p = lookup_path_to_sort_from env evd j.uj_type in let evd,j1 = apply_coercion env evd p j t in - let whd_evar evd c = EConstr.of_constr (whd_evar evd (EConstr.Unsafe.to_constr c)) in let j2 = on_judgment_type (whd_evar evd) j1 in (evd,type_judgment env evd j2) with Not_found | NoCoercion -> diff --git a/pretyping/evarconv.ml b/pretyping/evarconv.ml index 87267d538..3ae2e35e6 100644 --- a/pretyping/evarconv.ml +++ b/pretyping/evarconv.ml @@ -526,7 +526,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty | None, Success i' -> (* We do have sk1[] = sk2[]: we now unify ?ev1 and ?ev2 *) (* Note that ?ev1 and ?ev2, may have been instantiated in the meantime *) - let ev1' = EConstr.of_constr (whd_evar i' (EConstr.Unsafe.to_constr (mkEvar ev1))) in + let ev1' = whd_evar i' (mkEvar ev1) in if isEvar i' ev1' then solve_simple_eqn (evar_conv_x ts) env i' (position_problem true pbty,destEvar i' ev1', term2) @@ -536,7 +536,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty | Some (r,[]), Success i' -> (* We have sk1'[] = sk2[] for some sk1' s.t. sk1[]=sk1'[r[]] *) (* we now unify r[?ev1] and ?ev2 *) - let ev2' = EConstr.of_constr (whd_evar i' (EConstr.Unsafe.to_constr (mkEvar ev2))) in + let ev2' = whd_evar i' (mkEvar ev2) in if isEvar i' ev2' then solve_simple_eqn (evar_conv_x ts) env i' (position_problem false pbty,destEvar i' ev2',Stack.zip evd (term1,r)) @@ -547,7 +547,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty (* Symmetrically *) (* We have sk1[] = sk2'[] for some sk2' s.t. sk2[]=sk2'[r[]] *) (* we now unify ?ev1 and r[?ev2] *) - let ev1' = EConstr.of_constr (whd_evar i' (EConstr.Unsafe.to_constr (mkEvar ev1))) in + let ev1' = whd_evar i' (mkEvar ev1) in if isEvar i' ev1' then solve_simple_eqn (evar_conv_x ts) env i' (position_problem true pbty,destEvar i' ev1',Stack.zip evd (term2,r)) diff --git a/pretyping/evardefine.ml b/pretyping/evardefine.ml index 875e4a118..5831d3191 100644 --- a/pretyping/evardefine.ml +++ b/pretyping/evardefine.ml @@ -33,8 +33,9 @@ let new_evar_unsafe env evd ?src ?filter ?candidates ?store ?naming ?principal t (Sigma.to_evar_map evd, evk) let env_nf_evar sigma env = + let nf_evar c = EConstr.Unsafe.to_constr (nf_evar sigma (EConstr.of_constr c)) in process_rel_context - (fun d e -> push_rel (RelDecl.map_constr (nf_evar sigma) d) e) env + (fun d e -> push_rel (RelDecl.map_constr nf_evar d) e) env let env_nf_betaiotaevar sigma env = process_rel_context @@ -144,7 +145,7 @@ let define_pure_evar_as_lambda env evd evk = | _ -> error_not_product env evd typ in let avoid = ids_of_named_context (evar_context evi) in let id = - next_name_away_with_default_using_types "x" na avoid (Reductionops.whd_evar evd (EConstr.Unsafe.to_constr dom)) in + next_name_away_with_default_using_types "x" na avoid (EConstr.Unsafe.to_constr (Reductionops.whd_evar evd dom)) in let newenv = push_named (nlocal_assum (id, dom)) evenv in let filter = Filter.extend 1 (evar_filter evi) in let src = evar_source evk evd1 in diff --git a/pretyping/evarsolve.ml b/pretyping/evarsolve.ml index de2e46a78..3235c2505 100644 --- a/pretyping/evarsolve.ml +++ b/pretyping/evarsolve.ml @@ -601,13 +601,13 @@ let define_evar_from_virtual_equation define_fun env evd src t_in_env ty_t_in_si let evd = Sigma.Unsafe.of_evar_map evd in let Sigma (evar_in_env, evd, _) = new_evar_instance sign evd ty_t_in_sign ~filter ~src inst_in_env in let evd = Sigma.to_evar_map evd in - let t_in_env = EConstr.of_constr (whd_evar evd (EConstr.Unsafe.to_constr t_in_env)) in + let t_in_env = whd_evar evd t_in_env in let (evk, _) = destEvar evd evar_in_env in let evd = define_fun env evd None (destEvar evd evar_in_env) t_in_env in let ctxt = named_context_of_val sign in let inst_in_sign = inst_of_vars (Filter.filter_list filter ctxt) in let evar_in_sign = mkEvar (evk, inst_in_sign) in - (evd,EConstr.of_constr (whd_evar evd (EConstr.Unsafe.to_constr evar_in_sign))) + (evd,whd_evar evd evar_in_sign) (* We have x1..xq |- ?e1 : τ and had to solve something like * Σ; Γ |- ?e1[u1..uq] = (...\y1 ... \yk ... c), where c is typically some diff --git a/pretyping/pretype_errors.ml b/pretyping/pretype_errors.ml index 673554005..24f6d1689 100644 --- a/pretyping/pretype_errors.ml +++ b/pretyping/pretype_errors.ml @@ -10,50 +10,51 @@ open Util open Names open Term open Environ +open EConstr open Type_errors type unification_error = - | OccurCheck of existential_key * EConstr.constr - | NotClean of EConstr.existential * env * EConstr.constr (* Constr is a variable not in scope *) + | OccurCheck of existential_key * constr + | NotClean of existential * env * constr (* Constr is a variable not in scope *) | NotSameArgSize | NotSameHead | NoCanonicalStructure - | ConversionFailed of env * EConstr.constr * EConstr.constr (* Non convertible closed terms *) + | ConversionFailed of env * constr * constr (* Non convertible closed terms *) | MetaOccurInBody of existential_key - | InstanceNotSameType of existential_key * env * EConstr.types * EConstr.types + | InstanceNotSameType of existential_key * env * types * types | UnifUnivInconsistency of Univ.univ_inconsistency | CannotSolveConstraint of Evd.evar_constraint * unification_error | ProblemBeyondCapabilities type position = (Id.t * Locus.hyp_location_flag) option -type position_reporting = (position * int) * EConstr.t +type position_reporting = (position * int) * constr -type subterm_unification_error = bool * position_reporting * position_reporting * (EConstr.constr * EConstr.constr * unification_error) option +type subterm_unification_error = bool * position_reporting * position_reporting * (constr * constr * unification_error) option -type type_error = (EConstr.constr, EConstr.types) ptype_error +type type_error = (constr, types) ptype_error type pretype_error = (* Old Case *) - | CantFindCaseType of EConstr.constr + | CantFindCaseType of constr (* Type inference unification *) - | ActualTypeNotCoercible of EConstr.unsafe_judgment * EConstr.types * unification_error + | ActualTypeNotCoercible of unsafe_judgment * types * unification_error (* Tactic unification *) - | UnifOccurCheck of existential_key * EConstr.constr + | UnifOccurCheck of existential_key * constr | UnsolvableImplicit of existential_key * Evd.unsolvability_explanation option - | CannotUnify of EConstr.constr * EConstr.constr * unification_error option - | CannotUnifyLocal of EConstr.constr * EConstr.constr * EConstr.constr + | CannotUnify of constr * constr * unification_error option + | CannotUnifyLocal of constr * constr * constr | CannotUnifyBindingType of constr * constr | CannotGeneralize of constr - | NoOccurrenceFound of EConstr.constr * Id.t option - | CannotFindWellTypedAbstraction of EConstr.constr * EConstr.constr list * (env * type_error) option - | WrongAbstractionType of Name.t * EConstr.constr * EConstr.types * EConstr.types + | NoOccurrenceFound of constr * Id.t option + | CannotFindWellTypedAbstraction of constr * constr list * (env * type_error) option + | WrongAbstractionType of Name.t * constr * types * types | AbstractionOverMeta of Name.t * Name.t - | NonLinearUnification of Name.t * EConstr.constr + | NonLinearUnification of Name.t * constr (* Pretyping *) | VarNotFound of Id.t - | UnexpectedType of EConstr.constr * EConstr.constr - | NotProduct of EConstr.constr + | UnexpectedType of constr * constr + | NotProduct of constr | TypingError of type_error | CannotUnifyOccurrences of subterm_unification_error | UnsatisfiableConstraints of diff --git a/pretyping/pretype_errors.mli b/pretyping/pretype_errors.mli index 7cef14339..c303d5d94 100644 --- a/pretyping/pretype_errors.mli +++ b/pretyping/pretype_errors.mli @@ -45,8 +45,8 @@ type pretype_error = | UnsolvableImplicit of existential_key * Evd.unsolvability_explanation option | CannotUnify of constr * constr * unification_error option | CannotUnifyLocal of constr * constr * constr - | CannotUnifyBindingType of Constr.constr * Constr.constr - | CannotGeneralize of Constr.constr + | CannotUnifyBindingType of constr * constr + | CannotGeneralize of constr | NoOccurrenceFound of constr * Id.t option | CannotFindWellTypedAbstraction of constr * constr list * (env * type_error) option | WrongAbstractionType of Name.t * constr * types * types diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml index 6b6800ac6..4660978df 100644 --- a/pretyping/pretyping.ml +++ b/pretyping/pretyping.ml @@ -143,9 +143,6 @@ let nf_fix sigma (nas, cs, ts) = let inj c = EConstr.to_constr sigma c in (nas, Array.map inj cs, Array.map inj ts) -let nf_evar sigma c = - EConstr.of_constr (nf_evar sigma (EConstr.Unsafe.to_constr c)) - let search_guard loc env possible_indexes fixdefs = (* Standard situation with only one possibility for each fix. *) (* We treat it separately in order to get proper error msg. *) diff --git a/pretyping/reductionops.mli b/pretyping/reductionops.mli index 8aaeeb2c2..dcc11cfcf 100644 --- a/pretyping/reductionops.mli +++ b/pretyping/reductionops.mli @@ -156,11 +156,11 @@ val nf_beta : local_reduction_function val nf_betaiota : local_reduction_function val nf_betaiotazeta : local_reduction_function val nf_all : reduction_function -val nf_evar : evar_map -> Constr.constr -> Constr.constr +val nf_evar : evar_map -> constr -> constr (** Lazy strategy, weak head reduction *) -val whd_evar : evar_map -> Constr.constr -> Constr.constr +val whd_evar : evar_map -> constr -> constr val whd_nored : local_reduction_function val whd_beta : local_reduction_function val whd_betaiota : local_reduction_function diff --git a/pretyping/typing.ml b/pretyping/typing.ml index 7baff421f..e6f1e46b6 100644 --- a/pretyping/typing.ml +++ b/pretyping/typing.ml @@ -252,7 +252,7 @@ let judge_of_letin env name defj typj j = (* cstr must be in n.f. w.r.t. evars and execute returns a judgement where both the term and type are in n.f. *) let rec execute env evdref cstr = - let cstr = EConstr.of_constr (whd_evar !evdref (EConstr.Unsafe.to_constr cstr)) in + let cstr = whd_evar !evdref cstr in match EConstr.kind !evdref cstr with | Meta n -> { uj_val = cstr; uj_type = meta_type !evdref n } @@ -411,6 +411,6 @@ let e_solve_evars env evdref c = let env = enrich_env env evdref in let c = (execute env evdref c).uj_val in (* side-effect on evdref *) - EConstr.of_constr (nf_evar !evdref (EConstr.Unsafe.to_constr c)) + nf_evar !evdref c let _ = Evarconv.set_solve_evars (fun env evdref c -> e_solve_evars env evdref c) diff --git a/pretyping/unification.ml b/pretyping/unification.ml index 20f27a15a..1dc3ccdc0 100644 --- a/pretyping/unification.ml +++ b/pretyping/unification.ml @@ -137,7 +137,7 @@ let abstract_list_all_with_dependencies env evd typ c l = Evarconv.second_order_matching empty_transparent_state env evd ev' argoccs c in if b then - let p = nf_evar evd (EConstr.Unsafe.to_constr ev) in + let p = nf_evar evd ev in evd, p else error_cannot_find_well_typed_abstraction env evd c l None @@ -1240,7 +1240,7 @@ let try_to_coerce env evd c cty tycon = let j = make_judge c cty in let (evd',j') = inh_conv_coerce_rigid_to true Loc.ghost env evd j tycon in let evd' = Evarconv.consider_remaining_unif_problems env evd' in - let evd' = Evd.map_metas_fvalue (nf_evar evd') evd' in + let evd' = Evd.map_metas_fvalue (fun c -> EConstr.Unsafe.to_constr (nf_evar evd' (EConstr.of_constr c))) evd' in (evd',j'.uj_val) let w_coerce_to_type env evd c cty mvty = @@ -1397,7 +1397,7 @@ let w_merge env with_types flags (evd,metas,evars : subst0) = let evd''' = w_merge_rec evd'' mc ec [] in if evd' == evd''' then Evd.define sp (EConstr.Unsafe.to_constr c) evd''' - else Evd.define sp (Evarutil.nf_evar evd''' (EConstr.Unsafe.to_constr c)) evd''' in + else Evd.define sp (EConstr.Unsafe.to_constr (Evarutil.nf_evar evd''' c)) evd''' in let check_types evd = let metas = Evd.meta_list evd in @@ -1513,8 +1513,7 @@ let finish_evar_resolution ?(flags=Pretyping.all_and_fail_flags) env current_sig let current_sigma = Sigma.to_evar_map current_sigma in let sigma = Pretyping.solve_remaining_evars flags env current_sigma pending in let sigma, subst = nf_univ_variables sigma in - let c = EConstr.Unsafe.to_constr c in - Sigma.Unsafe.of_pair (EConstr.of_constr (CVars.subst_univs_constr subst (nf_evar sigma c)), sigma) + Sigma.Unsafe.of_pair (EConstr.of_constr (CVars.subst_univs_constr subst (EConstr.Unsafe.to_constr (nf_evar sigma c))), sigma) let default_matching_core_flags sigma = let ts = Names.full_transparent_state in { @@ -1602,7 +1601,7 @@ let make_pattern_test from_prefix_of_ind is_correct_type env sigma (pending,c) = (fun test -> match test.testing_state with | None -> None | Some (sigma,_,l) -> - let c = applist (EConstr.of_constr (nf_evar sigma (EConstr.Unsafe.to_constr (local_strong whd_meta sigma c))), l) in + let c = applist (nf_evar sigma (local_strong whd_meta sigma c), l) in let univs, subst = nf_univ_variables sigma in Some (sigma,EConstr.of_constr (CVars.subst_univs_constr subst (EConstr.Unsafe.to_constr c)))) @@ -1926,7 +1925,6 @@ let secondOrderAbstraction env evd flags typ (p, oplist) = let secondOrderDependentAbstraction env evd flags typ (p, oplist) = let typp = Typing.meta_type evd p in let evd, pred = abstract_list_all_with_dependencies env evd typp typ oplist in - let pred = EConstr.of_constr pred in w_merge env false flags.merge_unify_flags (evd,[p,pred,(Conv,TypeProcessed)],[]) diff --git a/printing/printer.ml b/printing/printer.ml index ba4b68296..63aeec82c 100644 --- a/printing/printer.ml +++ b/printing/printer.ml @@ -506,8 +506,8 @@ let print_evar_constraints gl sigma = end in let pr_evconstr (pbty,env,t1,t2) = - let t1 = Evarutil.nf_evar sigma t1 - and t2 = Evarutil.nf_evar sigma t2 in + let t1 = Evarutil.nf_evar sigma (EConstr.of_constr t1) + and t2 = Evarutil.nf_evar sigma (EConstr.of_constr t2) in 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 @@ -517,11 +517,11 @@ let print_evar_constraints gl sigma = naming/renaming *) Namegen.make_all_name_different env in str" " ++ - hov 2 (pr_env env ++ pr_lconstr_env env sigma t1 ++ spc () ++ + hov 2 (pr_env env ++ pr_leconstr_env env sigma t1 ++ spc () ++ str (match pbty with | Reduction.CONV -> "==" | Reduction.CUMUL -> "<=") ++ - spc () ++ pr_lconstr_env env sigma t2) + spc () ++ pr_leconstr_env env sigma t2) in let pr_candidate ev evi (candidates,acc) = if Option.has_some evi.evar_candidates then diff --git a/proofs/pfedit.ml b/proofs/pfedit.ml index af910810b..142fd9a89 100644 --- a/proofs/pfedit.ml +++ b/proofs/pfedit.ml @@ -199,6 +199,7 @@ let refine_by_tactic env sigma ty tac = | _ -> assert false in let ans = Reductionops.nf_evar sigma ans in + let ans = EConstr.Unsafe.to_constr ans in (** [neff] contains the freshly generated side-effects *) let neff = Evd.eval_side_effects sigma in (** Reset the old side-effects *) @@ -232,7 +233,8 @@ let solve_by_implicit_tactic env sigma evk = let tac = Proofview.tclTHEN tac (Proofview.tclEXTEND [] (Proofview.tclZERO (CErrors.UserError (None,Pp.str"Proof is not complete."))) []) in (try let c = Evarutil.nf_evars_universes sigma evi.evar_concl in - if Evarutil.has_undefined_evars sigma (EConstr.of_constr c) then raise Exit; + let c = EConstr.of_constr c in + if Evarutil.has_undefined_evars sigma c then raise Exit; let (ans, _, ctx) = build_by_tactic env (Evd.evar_universe_context sigma) c tac in let sigma = Evd.set_universe_context sigma ctx in diff --git a/proofs/pfedit.mli b/proofs/pfedit.mli index 807a554df..516125ea1 100644 --- a/proofs/pfedit.mli +++ b/proofs/pfedit.mli @@ -60,7 +60,7 @@ type lemma_possible_guards = Proof_global.lemma_possible_guards type universe_binders = Id.t Loc.located list val start_proof : - Id.t -> ?pl:universe_binders -> goal_kind -> Evd.evar_map -> named_context_val -> constr -> + Id.t -> ?pl:universe_binders -> goal_kind -> Evd.evar_map -> named_context_val -> EConstr.constr -> ?init_tac:unit Proofview.tactic -> Proof_global.proof_terminator -> unit @@ -106,7 +106,7 @@ val get_current_context : unit -> Evd.evar_map * env (** [current_proof_statement] *) val current_proof_statement : - unit -> Id.t * goal_kind * types + unit -> Id.t * goal_kind * EConstr.types (** {6 ... } *) (** [get_current_proof_name ()] return the name of the current focused @@ -166,15 +166,15 @@ val instantiate_nth_evar_com : int -> Constrexpr.constr_expr -> unit val build_constant_by_tactic : Id.t -> Evd.evar_universe_context -> named_context_val -> ?goal_kind:goal_kind -> - types -> unit Proofview.tactic -> + EConstr.types -> unit Proofview.tactic -> Safe_typing.private_constants Entries.definition_entry * bool * Evd.evar_universe_context val build_by_tactic : ?side_eff:bool -> env -> Evd.evar_universe_context -> ?poly:polymorphic -> - types -> unit Proofview.tactic -> + EConstr.types -> unit Proofview.tactic -> constr * bool * Evd.evar_universe_context -val refine_by_tactic : env -> Evd.evar_map -> types -> unit Proofview.tactic -> +val refine_by_tactic : env -> Evd.evar_map -> EConstr.types -> unit Proofview.tactic -> constr * Evd.evar_map (** A variant of the above function that handles open terms as well. Caveat: all effects are purged in the returned term at the end, but other diff --git a/proofs/proof.mli b/proofs/proof.mli index 5053fc7fb..8dcc5b76e 100644 --- a/proofs/proof.mli +++ b/proofs/proof.mli @@ -66,9 +66,9 @@ val map_structured_proof : proof -> (Evd.evar_map -> Goal.goal -> 'a) -> ('a pre (*** General proof functions ***) -val start : Evd.evar_map -> (Environ.env * Term.types) list -> proof +val start : Evd.evar_map -> (Environ.env * EConstr.types) list -> proof val dependent_start : Proofview.telescope -> proof -val initial_goals : proof -> (Term.constr * Term.types) list +val initial_goals : proof -> (EConstr.constr * EConstr.types) list val initial_euctx : proof -> Evd.evar_universe_context (* Returns [true] if the considered proof is completed, that is if no goal remain @@ -79,7 +79,7 @@ val is_done : proof -> bool val is_complete : proof -> bool (* Returns the list of partial proofs to initial goals. *) -val partial_proof : proof -> Term.constr list +val partial_proof : proof -> EConstr.constr list val compact : proof -> proof diff --git a/proofs/proof_global.ml b/proofs/proof_global.ml index 2956d623f..eb1bea897 100644 --- a/proofs/proof_global.ml +++ b/proofs/proof_global.ml @@ -375,6 +375,7 @@ let close_proof ~keep_body_ucst_separate ?feedback_id ~now fpl = in let entries = Future.map2 (fun p (_, t) -> + let t = EConstr.Unsafe.to_constr t in let univstyp, body = make_body t p in let univs, typ = Future.force univstyp in { Entries. @@ -406,7 +407,7 @@ let return_proof ?(allow_partial=false) () = (** ppedrot: FIXME, this is surely wrong. There is no reason to duplicate side-effects... This may explain why one need to uniquize side-effects thereafter... *) - let proofs = List.map (fun c -> c, eff) proofs in + let proofs = List.map (fun c -> EConstr.Unsafe.to_constr c, eff) proofs in proofs, Evd.evar_universe_context evd end else let initial_goals = Proof.initial_goals proof in @@ -430,7 +431,7 @@ let return_proof ?(allow_partial=false) () = side-effects... This may explain why one need to uniquize side-effects thereafter... *) let proofs = - List.map (fun (c, _) -> (Evarutil.nf_evars_universes evd c, eff)) initial_goals in + List.map (fun (c, _) -> (Evarutil.nf_evars_universes evd (EConstr.Unsafe.to_constr c), eff)) initial_goals in proofs, Evd.evar_universe_context evd let close_future_proof ~feedback_id proof = diff --git a/proofs/proof_global.mli b/proofs/proof_global.mli index 97a21cf22..3b2cc6b23 100644 --- a/proofs/proof_global.mli +++ b/proofs/proof_global.mli @@ -86,7 +86,7 @@ val apply_terminator : proof_terminator -> proof_ending -> unit end of the proof to close the proof. *) val start_proof : Evd.evar_map -> Names.Id.t -> ?pl:universe_binders -> - Decl_kinds.goal_kind -> (Environ.env * Term.types) list -> + Decl_kinds.goal_kind -> (Environ.env * EConstr.types) list -> proof_terminator -> unit (** Like [start_proof] except that there may be dependencies between @@ -201,7 +201,7 @@ end val get_default_goal_selector : unit -> Vernacexpr.goal_selector module V82 : sig - val get_current_initial_conclusions : unit -> Names.Id.t *(Term.types list * + val get_current_initial_conclusions : unit -> Names.Id.t *(EConstr.types list * Decl_kinds.goal_kind) end diff --git a/proofs/tacmach.ml b/proofs/tacmach.ml index 2fab026ea..3ed262d6e 100644 --- a/proofs/tacmach.ml +++ b/proofs/tacmach.ml @@ -223,7 +223,7 @@ module New = struct let gl = Proofview.Goal.assume gl in let concl = Proofview.Goal.concl gl in let sigma = project gl in - EConstr.of_constr (nf_evar sigma (EConstr.Unsafe.to_constr concl)) + nf_evar sigma concl let pf_whd_all gl t = pf_apply whd_all gl t @@ -241,6 +241,6 @@ module New = struct let pf_whd_all gl t = pf_apply whd_all gl t let pf_compute gl t = pf_apply compute gl t - let pf_nf_evar gl t = EConstr.of_constr (nf_evar (project gl) (EConstr.Unsafe.to_constr t)) + let pf_nf_evar gl t = nf_evar (project gl) t end diff --git a/stm/lemmas.ml b/stm/lemmas.ml index e3f1c1ac2..77fb960e1 100644 --- a/stm/lemmas.ml +++ b/stm/lemmas.ml @@ -449,7 +449,7 @@ let start_proof_with_initialization kind ctx recguard thms snl hook = List.iter (fun (strength,ref,imps) -> maybe_declare_manual_implicits false ref imps; call_hook (fun exn -> exn) hook strength ref) thms_data in - start_proof_univs id ?pl kind ctx t ?init_tac (fun ctx -> mk_hook (hook ctx)) ~compute_guard:guard + start_proof_univs id ?pl kind ctx (EConstr.of_constr t) ?init_tac (fun ctx -> mk_hook (hook ctx)) ~compute_guard:guard let start_proof_com ?inference_hook kind thms hook = let env0 = Global.env () in @@ -461,13 +461,12 @@ let start_proof_com ?inference_hook kind thms hook = let thms = List.map (fun (sopt,(bl,t,guard)) -> let impls, ((env, ctx), imps) = interp_context_evars env0 evdref bl in let t', imps' = interp_type_evars_impls ~impls env evdref t in - let t' = EConstr.Unsafe.to_constr t' in let flags = all_and_fail_flags in let flags = { flags with use_hook = inference_hook } in evdref := solve_remaining_evars flags env !evdref (Evd.empty,!evdref); let ids = List.map RelDecl.get_name ctx in (compute_proof_name (pi1 kind) sopt, - (nf_evar !evdref (Term.it_mkProd_or_LetIn t' ctx), + (EConstr.Unsafe.to_constr (nf_evar !evdref (EConstr.it_mkProd_or_LetIn t' ctx)), (ids, imps @ lift_implicits (List.length ids) imps'), guard))) thms in @@ -519,6 +518,7 @@ let save_proof ?proof = function | None -> let pftree = Pfedit.get_pftreestate () in let id, k, typ = Pfedit.current_proof_statement () in + let typ = EConstr.Unsafe.to_constr typ in let universes = Proof.initial_euctx pftree in (* This will warn if the proof is complete *) let pproofs, _univs = diff --git a/stm/lemmas.mli b/stm/lemmas.mli index 39c089be9..681413a29 100644 --- a/stm/lemmas.mli +++ b/stm/lemmas.mli @@ -19,17 +19,17 @@ val call_hook : Future.fix_exn -> 'a declaration_hook -> Decl_kinds.locality -> Globnames.global_reference -> 'a (** A hook start_proof calls on the type of the definition being started *) -val set_start_hook : (types -> unit) -> unit +val set_start_hook : (EConstr.types -> unit) -> unit val start_proof : Id.t -> ?pl:universe_binders -> goal_kind -> Evd.evar_map -> ?terminator:(lemma_possible_guards -> unit declaration_hook -> Proof_global.proof_terminator) -> - ?sign:Environ.named_context_val -> types -> + ?sign:Environ.named_context_val -> EConstr.types -> ?init_tac:unit Proofview.tactic -> ?compute_guard:lemma_possible_guards -> unit declaration_hook -> unit val start_proof_univs : Id.t -> ?pl:universe_binders -> goal_kind -> Evd.evar_map -> ?terminator:(lemma_possible_guards -> (Evd.evar_universe_context option -> unit declaration_hook) -> Proof_global.proof_terminator) -> - ?sign:Environ.named_context_val -> types -> + ?sign:Environ.named_context_val -> EConstr.types -> ?init_tac:unit Proofview.tactic -> ?compute_guard:lemma_possible_guards -> (Evd.evar_universe_context option -> unit declaration_hook) -> unit diff --git a/stm/stm.ml b/stm/stm.ml index d60412c0c..47e806929 100644 --- a/stm/stm.ml +++ b/stm/stm.ml @@ -1719,8 +1719,10 @@ end = struct (* {{{ *) match Evd.(evar_body (find sigma r_goal)) with | Evd.Evar_empty -> RespNoProgress | Evd.Evar_defined t -> + let t = EConstr.of_constr t in let t = Evarutil.nf_evar sigma t in - if Evarutil.is_ground_term sigma (EConstr.of_constr t) then + if Evarutil.is_ground_term sigma t then + let t = EConstr.Unsafe.to_constr t in RespBuiltSubProof (t, Evd.evar_universe_context sigma) else CErrors.user_err ~hdr:"STM" (str"The solution is not ground") end) () diff --git a/tactics/auto.ml b/tactics/auto.ml index b548f8b92..17a488ddb 100644 --- a/tactics/auto.ml +++ b/tactics/auto.ml @@ -321,7 +321,7 @@ let rec trivial_fail_db dbg mod_delta db_list local_db = ( Proofview.Goal.enter { enter = begin fun gl -> let sigma = Tacmach.New.project gl in let env = Proofview.Goal.env gl in - let nf c = Evarutil.nf_evar sigma c in + let nf c = EConstr.Unsafe.to_constr (Evarutil.nf_evar sigma (EConstr.of_constr c)) in let decl = Tacmach.New.pf_last_hyp (Proofview.Goal.assume gl) in let hyp = Context.Named.Declaration.map_constr nf decl in let hintl = make_resolve_hyp env sigma hyp diff --git a/tactics/class_tactics.ml b/tactics/class_tactics.ml index fa2c21ac3..a4b6cb53b 100644 --- a/tactics/class_tactics.ml +++ b/tactics/class_tactics.ml @@ -521,8 +521,7 @@ let evars_to_goals p evm = (** Making local hints *) let make_resolve_hyp env sigma st flags only_classes pri decl = let id = NamedDecl.get_id decl in - let cty = Evarutil.nf_evar sigma (NamedDecl.get_type decl) in - let cty = EConstr.of_constr cty in + let cty = Evarutil.nf_evar sigma (EConstr.of_constr (NamedDecl.get_type decl)) in let rec iscl env ty = let ctx, ar = decompose_prod_assum sigma ty in match EConstr.kind sigma (fst (decompose_app sigma ar)) with @@ -1476,8 +1475,7 @@ let resolve_one_typeclass env ?(sigma=Evd.empty) gl unique = in let evd = sig_sig gls' in let t' = mkEvar (ev, Array.of_list subst) in - let term = Evarutil.nf_evar evd (EConstr.Unsafe.to_constr t') in - let term = EConstr.of_constr term in + let term = Evarutil.nf_evar evd t' in evd, term let _ = diff --git a/tactics/contradiction.ml b/tactics/contradiction.ml index a3a448aad..7173fb4fd 100644 --- a/tactics/contradiction.ml +++ b/tactics/contradiction.ml @@ -66,8 +66,7 @@ let contradiction_context = | [] -> Tacticals.New.tclZEROMSG (Pp.str"No such contradiction") | d :: rest -> let id = NamedDecl.get_id d in - let typ = nf_evar sigma (NamedDecl.get_type d) in - let typ = EConstr.of_constr typ in + let typ = nf_evar sigma (EConstr.of_constr (NamedDecl.get_type d)) in let typ = whd_all env sigma typ in if is_empty_type sigma typ then simplest_elim (mkVar id) diff --git a/tactics/eauto.ml b/tactics/eauto.ml index 01f21910c..7453fff5c 100644 --- a/tactics/eauto.ml +++ b/tactics/eauto.ml @@ -289,8 +289,7 @@ module SearchProblem = struct in let rec_tacs = let l = - let concl = Reductionops.nf_evar (project g) (EConstr.Unsafe.to_constr (pf_concl g)) in - let concl = EConstr.of_constr concl in + let concl = Reductionops.nf_evar (project g) (pf_concl g) in filter_tactics s.tacres (e_possible_resolve (project g) s.dblist (List.hd s.localdb) secvars concl) in diff --git a/tactics/equality.ml b/tactics/equality.ml index c80cf4416..072da995d 100644 --- a/tactics/equality.ml +++ b/tactics/equality.ml @@ -1212,7 +1212,7 @@ let sig_clausal_form env sigma sort_of_ty siglen ty dflt = error "Cannot solve a unification problem." in let scf = sigrec_clausal_form siglen ty in - !evdref, EConstr.of_constr (Evarutil.nf_evar !evdref (EConstr.Unsafe.to_constr scf)) + !evdref, Evarutil.nf_evar !evdref scf (* The problem is to build a destructor (a generalization of the predecessor) which, when applied to a term made of constructors diff --git a/tactics/hints.ml b/tactics/hints.ml index 851e9f01f..ef97b0b33 100644 --- a/tactics/hints.ml +++ b/tactics/hints.ml @@ -1200,7 +1200,8 @@ let prepare_hint check (poly,local) env init (sigma,c) = It is actually a bit stupid to generalize over evars since the first thing make_resolves will do is to re-instantiate the products *) let sigma, subst = Evd.nf_univ_variables sigma in - let c = Evarutil.nf_evar sigma (EConstr.Unsafe.to_constr c) in + let c = Evarutil.nf_evar sigma c in + let c = EConstr.Unsafe.to_constr c in let c = CVars.subst_univs_constr subst c in let c = EConstr.of_constr c in let c = drop_extra_implicit_args sigma c in diff --git a/tactics/leminv.ml b/tactics/leminv.ml index 3199623e7..a05b4fbf3 100644 --- a/tactics/leminv.ml +++ b/tactics/leminv.ml @@ -208,14 +208,12 @@ let inversion_scheme env sigma t sort dep_option inv_op = user_err ~hdr:"lemma_inversion" (str"Computed inversion goal was not closed in initial signature."); *) - let invGoal = EConstr.Unsafe.to_constr invGoal in let pf = Proof.start (Evd.from_ctx (evar_universe_context sigma)) [invEnv,invGoal] in let pf = fst (Proof.run_tactic env ( tclTHEN intro (onLastHypId inv_op)) pf) in let pfterm = List.hd (Proof.partial_proof pf) in - let pfterm = EConstr.of_constr pfterm in let global_named_context = Global.named_context_val () in let ownSign = ref begin fold_named_context diff --git a/tactics/tactics.ml b/tactics/tactics.ml index 10582288c..0ecccd5c0 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -2945,7 +2945,6 @@ let quantify lconstr = (* Modifying/Adding an hypothesis *) let specialize (c,lbind) ipat = - let nf_evar sigma c = EConstr.of_constr (nf_evar sigma (EConstr.Unsafe.to_constr c)) in Proofview.Goal.enter { enter = begin fun gl -> let env = Proofview.Goal.env gl in let sigma = Sigma.to_evar_map (Proofview.Goal.sigma gl) in @@ -3825,7 +3824,7 @@ let specialize_eqs id gl = let acc' = it_mkLambda_or_LetIn acc ctx'' in let ty' = Tacred.whd_simpl env !evars ty' and acc' = Tacred.whd_simpl env !evars acc' in - let ty' = EConstr.of_constr (Evarutil.nf_evar !evars (EConstr.Unsafe.to_constr ty')) in + let ty' = Evarutil.nf_evar !evars ty' in if worked then tclTHENFIRST (Tacmach.internal_cut true id ty') (Proofview.V82.of_tactic (exact_no_check ((* refresh_universes_strict *) acc'))) gl @@ -4951,6 +4950,7 @@ let abstract_subproof id gk tac = let ctx = Evd.universe_context_set evd in evd, ctx, nf concl in + let concl = EConstr.of_constr concl in let solve_tac = tclCOMPLETE (tclTHEN (tclDO (List.length sign) intro) tac) in let ectx = Evd.evar_universe_context evd in let (const, safe, ectx) = diff --git a/toplevel/auto_ind_decl.ml b/toplevel/auto_ind_decl.ml index 5e686c41e..14071d869 100644 --- a/toplevel/auto_ind_decl.ml +++ b/toplevel/auto_ind_decl.ml @@ -658,6 +658,7 @@ let make_bl_scheme mode mind = let bl_goal, eff = compute_bl_goal ind lnamesparrec nparrec in let ctx = Evd.make_evar_universe_context (Global.env ()) None in let side_eff = side_effect_of_mode mode in + let bl_goal = EConstr.of_constr bl_goal in let (ans, _, ctx) = Pfedit.build_by_tactic ~side_eff (Global.env()) ctx bl_goal (compute_bl_tact mode (!bl_scheme_kind_aux()) (ind, Univ.Instance.empty) lnamesparrec nparrec) in @@ -783,6 +784,7 @@ let make_lb_scheme mode mind = let lb_goal, eff = compute_lb_goal ind lnamesparrec nparrec in let ctx = Evd.make_evar_universe_context (Global.env ()) None in let side_eff = side_effect_of_mode mode in + let lb_goal = EConstr.of_constr lb_goal in let (ans, _, ctx) = Pfedit.build_by_tactic ~side_eff (Global.env()) ctx lb_goal (compute_lb_tact mode (!lb_scheme_kind_aux()) ind lnamesparrec nparrec) in @@ -956,7 +958,7 @@ let make_eq_decidability mode mind = context_chop (nparams-nparrec) mib.mind_params_ctxt in let side_eff = side_effect_of_mode mode in let (ans, _, ctx) = Pfedit.build_by_tactic ~side_eff (Global.env()) ctx - (compute_dec_goal (ind,u) lnamesparrec nparrec) + (EConstr.of_constr (compute_dec_goal (ind,u) lnamesparrec nparrec)) (compute_dec_tact ind lnamesparrec nparrec) in ([|ans|], ctx), Safe_typing.empty_private_constants diff --git a/toplevel/classes.ml b/toplevel/classes.ml index 5087aa0c8..7a5d0ed81 100644 --- a/toplevel/classes.ml +++ b/toplevel/classes.ml @@ -181,7 +181,7 @@ let new_instance ?(abstract=false) ?(global=false) ?(refine= !refine_instance) p let env' = push_rel_context ctx env in evars := Evarutil.nf_evar_map !evars; evars := resolve_typeclasses ~filter:Typeclasses.all_evars ~fail:true env !evars; - let subst = List.map (Evarutil.nf_evar !evars) subst in + let subst = List.map (fun c -> EConstr.Unsafe.to_constr (Evarutil.nf_evar !evars (EConstr.of_constr c))) subst in if abstract then begin let subst = List.fold_left2 @@ -327,7 +327,7 @@ let new_instance ?(abstract=false) ?(global=false) ?(refine= !refine_instance) p the refinement manually.*) let gls = List.rev (Evd.future_goals evm) in let evm = Evd.reset_future_goals evm in - Lemmas.start_proof id kind evm termtype + Lemmas.start_proof id kind evm (EConstr.of_constr termtype) (Lemmas.mk_hook (fun _ -> instance_hook k pri global imps ?hook)); (* spiwack: I don't know what to do with the status here. *) diff --git a/toplevel/command.ml b/toplevel/command.ml index 7e3828131..44fc4eb1a 100644 --- a/toplevel/command.ml +++ b/toplevel/command.ml @@ -306,7 +306,8 @@ let do_assumptions_unbound_univs (_, poly, _ as kind) nl l = (env,empty_internalization_env) l in let evd = solve_remaining_evars all_and_fail_flags env !evdref (Evd.empty,!evdref) in - let l = List.map (on_pi2 (nf_evar evd)) l in + let nf_evar c = EConstr.Unsafe.to_constr (nf_evar evd (EConstr.of_constr c)) in + let l = List.map (on_pi2 nf_evar) l in snd (List.fold_left (fun (subst,status) ((is_coe,idl),t,(ctx,imps)) -> let t = replace_vars subst t in let (refs,status') = declare_assumptions idl is_coe kind (t,ctx) [] imps false nl in @@ -465,7 +466,7 @@ let sign_level env evd sign = | LocalDef _ -> lev, push_rel d env | LocalAssum _ -> let s = destSort (Reduction.whd_all env - (nf_evar evd (EConstr.Unsafe.to_constr (Retyping.get_type_of env evd (EConstr.of_constr (RelDecl.get_type d)))))) + (EConstr.Unsafe.to_constr (nf_evar evd (Retyping.get_type_of env evd (EConstr.of_constr (RelDecl.get_type d)))))) in let u = univ_of_sort s in (Univ.sup u lev, push_rel d env)) @@ -905,8 +906,9 @@ let fix_sub_ref = make_ref fixsub_module "Fix_sub" let measure_on_R_ref = make_ref fixsub_module "MR" let well_founded = init_constant ["Init"; "Wf"] "well_founded" let mkSubset name typ prop = - mkApp (Universes.constr_of_global (delayed_force build_sigma).typ, - [| typ; mkLambda (name, typ, prop) |]) + let open EConstr in + EConstr.Unsafe.to_constr (mkApp (EConstr.of_constr (Universes.constr_of_global (delayed_force build_sigma).typ), + [| typ; mkLambda (name, typ, prop) |])) let sigT = Lazy.from_fun build_sigma_type let make_qref s = Qualid (Loc.ghost, qualid_of_string s) @@ -943,9 +945,11 @@ let rec telescope = function ty, (LocalDef (n, b, t) :: subst), lift 1 term let nf_evar_context sigma ctx = - List.map (map_constr (Evarutil.nf_evar sigma)) ctx + List.map (map_constr (fun c -> EConstr.Unsafe.to_constr (Evarutil.nf_evar sigma (EConstr.of_constr c)))) ctx let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation = + let open EConstr in + let open Vars in Coqlib.check_required_library ["Coq";"Program";"Wf"]; let env = Global.env() in let ctx = Evd.make_evar_universe_context env pl in @@ -954,11 +958,12 @@ let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation = let len = List.length binders_rel in let top_env = push_rel_context binders_rel env in let top_arity = interp_type_evars top_env evdref arityc in - let top_arity = EConstr.Unsafe.to_constr top_arity in - let full_arity = Term.it_mkProd_or_LetIn top_arity binders_rel in + let full_arity = it_mkProd_or_LetIn top_arity binders_rel in let argtyp, letbinders, make = telescope binders_rel in + let make = EConstr.of_constr make in let argname = Id.of_string "recarg" in let arg = LocalAssum (Name argname, argtyp) in + let argtyp = EConstr.of_constr argtyp in let binders = letbinders @ [arg] in let binders_env = push_rel_context binders_rel env in let rel, _ = interp_constr_evars_impls env evdref r in @@ -977,22 +982,21 @@ let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation = | _, _ -> error () with e when CErrors.noncritical e -> error () in - let rel = EConstr.Unsafe.to_constr rel in let measure = interp_casted_constr_evars binders_env evdref measure relargty in - let measure = EConstr.Unsafe.to_constr measure in let wf_rel, wf_rel_fun, measure_fn = let measure_body, measure = it_mkLambda_or_LetIn measure letbinders, it_mkLambda_or_LetIn measure binders in - let comb = Universes.constr_of_global (delayed_force measure_on_R_ref) in + let comb = EConstr.of_constr (Universes.constr_of_global (delayed_force measure_on_R_ref)) in + let relargty = EConstr.of_constr relargty in let wf_rel = mkApp (comb, [| argtyp; relargty; rel; measure |]) in let wf_rel_fun x y = mkApp (rel, [| subst1 x measure_body; subst1 y measure_body |]) in wf_rel, wf_rel_fun, measure in - let wf_proof = mkApp (delayed_force well_founded, [| argtyp ; wf_rel |]) in + let wf_proof = mkApp (EConstr.of_constr (delayed_force well_founded), [| argtyp ; wf_rel |]) in let argid' = Id.of_string (Id.to_string argname ^ "'") in let wfarg len = LocalAssum (Name argid', mkSubset (Name argid') argtyp @@ -1000,7 +1004,7 @@ let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation = in let intern_bl = wfarg 1 :: [arg] in let _intern_env = push_rel_context intern_bl env in - let proj = (*FIXME*)Universes.constr_of_global (delayed_force build_sigma).Coqlib.proj1 in + let proj = (*FIXME*)EConstr.of_constr (Universes.constr_of_global (delayed_force build_sigma).Coqlib.proj1) in let wfargpred = mkLambda (Name argid', argtyp, wf_rel_fun (mkRel 1) (mkRel 3)) in let projection = (* in wfarg :: arg :: before *) mkApp (proj, [| argtyp ; wfargpred ; mkRel 1 |]) @@ -1009,17 +1013,21 @@ let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation = let intern_arity = substl [projection] top_arity_let in (* substitute the projection of wfarg for something, now intern_arity is in wfarg :: arg *) - let intern_fun_arity_prod = Term.it_mkProd_or_LetIn intern_arity [wfarg 1] in + let intern_fun_arity_prod = it_mkProd_or_LetIn intern_arity [wfarg 1] in + let intern_fun_arity_prod = EConstr.Unsafe.to_constr intern_fun_arity_prod in let intern_fun_binder = LocalAssum (Name (add_suffix recname "'"), intern_fun_arity_prod) in let curry_fun = let wfpred = mkLambda (Name argid', argtyp, wf_rel_fun (mkRel 1) (mkRel (2 * len + 4))) in - let intro = (*FIXME*)Universes.constr_of_global (delayed_force build_sigma).Coqlib.intro in + let intro = (*FIXME*)EConstr.of_constr (Universes.constr_of_global (delayed_force build_sigma).Coqlib.intro) in let arg = mkApp (intro, [| argtyp; wfpred; lift 1 make; mkRel 1 |]) in let app = mkApp (mkRel (2 * len + 2 (* recproof + orig binders + current binders *)), [| arg |]) in let rcurry = mkApp (rel, [| measure; lift len measure |]) in + let rcurry = EConstr.Unsafe.to_constr rcurry in let lam = LocalAssum (Name (Id.of_string "recproof"), rcurry) in let body = it_mkLambda_or_LetIn app (lam :: binders_rel) in - let ty = Term.it_mkProd_or_LetIn (lift 1 top_arity) (lam :: binders_rel) in + let ty = it_mkProd_or_LetIn (lift 1 top_arity) (lam :: binders_rel) in + let body = EConstr.Unsafe.to_constr body in + let ty = EConstr.Unsafe.to_constr ty in LocalDef (Name recname, body, ty) in let fun_bl = intern_fun_binder :: [arg] in @@ -1028,26 +1036,24 @@ let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation = let ctx = LocalAssum (Name recname, get_type curry_fun) :: binders_rel in let (r, l, impls, scopes) = Constrintern.compute_internalization_data env - Constrintern.Recursive full_arity impls + Constrintern.Recursive (EConstr.Unsafe.to_constr full_arity) impls in let newimpls = Id.Map.singleton recname (r, l, impls @ [(Some (Id.of_string "recproof", Impargs.Manual, (true, false)))], scopes @ [None]) in interp_casted_constr_evars (push_rel_context ctx env) evdref - ~impls:newimpls body (lift 1 top_arity) + ~impls:newimpls body (EConstr.Unsafe.to_constr (lift 1 top_arity)) in - let intern_body = EConstr.Unsafe.to_constr intern_body in let intern_body_lam = it_mkLambda_or_LetIn intern_body (curry_fun :: lift_lets @ fun_bl) in let prop = mkLambda (Name argname, argtyp, top_arity_let) in let def = - mkApp (Universes.constr_of_global (delayed_force fix_sub_ref), + mkApp (EConstr.of_constr (Universes.constr_of_global (delayed_force fix_sub_ref)), [| argtyp ; wf_rel ; - EConstr.Unsafe.to_constr (Evarutil.e_new_evar env evdref - ~src:(Loc.ghost, Evar_kinds.QuestionMark (Evar_kinds.Define false)) (EConstr.of_constr wf_proof)); + Evarutil.e_new_evar env evdref + ~src:(Loc.ghost, Evar_kinds.QuestionMark (Evar_kinds.Define false)) wf_proof; prop |]) in - let def = Typing.e_solve_evars env evdref (EConstr.of_constr def) in - let def = EConstr.Unsafe.to_constr def in + let def = Typing.e_solve_evars env evdref def in let _ = evdref := Evarutil.nf_evar_map !evdref in let def = mkApp (def, [|intern_body_lam|]) in let binders_rel = nf_evar_context !evdref binders_rel in @@ -1057,21 +1063,22 @@ let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation = if List.length binders_rel > 1 then let name = add_suffix recname "_func" in let hook l gr _ = - let body = it_mkLambda_or_LetIn (mkApp (Universes.constr_of_global gr, [|make|])) binders_rel in - let ty = Term.it_mkProd_or_LetIn top_arity binders_rel in + let body = it_mkLambda_or_LetIn (mkApp (EConstr.of_constr (Universes.constr_of_global gr), [|make|])) binders_rel in + let ty = it_mkProd_or_LetIn top_arity binders_rel in + let ty = EConstr.Unsafe.to_constr ty in let pl, univs = Evd.universe_context ?names:pl !evdref in (*FIXME poly? *) - let ce = definition_entry ~poly ~types:ty ~univs (Evarutil.nf_evar !evdref body) in + let ce = definition_entry ~poly ~types:ty ~univs (EConstr.Unsafe.to_constr (Evarutil.nf_evar !evdref body)) in (** FIXME: include locality *) let c = Declare.declare_constant recname (DefinitionEntry ce, IsDefinition Definition) in let gr = ConstRef c in if Impargs.is_implicit_args () || not (List.is_empty impls) then Impargs.declare_manual_implicits false gr [impls] in - let typ = Term.it_mkProd_or_LetIn top_arity binders in + let typ = it_mkProd_or_LetIn top_arity binders in hook, name, typ else - let typ = Term.it_mkProd_or_LetIn top_arity binders_rel in + let typ = it_mkProd_or_LetIn top_arity binders_rel in let hook l gr _ = if Impargs.is_implicit_args () || not (List.is_empty impls) then Impargs.declare_manual_implicits false gr [impls] @@ -1080,6 +1087,8 @@ let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation = let hook = Lemmas.mk_hook hook in let fullcoqc = Evarutil.nf_evar !evdref def in let fullctyp = Evarutil.nf_evar !evdref typ in + let fullcoqc = EConstr.Unsafe.to_constr fullcoqc in + let fullctyp = EConstr.Unsafe.to_constr fullctyp in Obligations.check_evars env !evdref; let evars, _, evars_def, evars_typ = Obligations.eterm_obligations env recname !evdref 0 fullcoqc fullctyp @@ -1110,7 +1119,7 @@ let interp_recursive isfix fixl notations = let fixctximpenvs, fixctximps = List.split fiximppairs in let fixccls,fixcclimps = List.split (List.map3 (interp_fix_ccl evdref) fixctximpenvs fixctxs fixl) in let fixtypes = List.map2 build_fix_type fixctxs fixccls in - let fixtypes = List.map (nf_evar !evdref) fixtypes in + let fixtypes = List.map (fun c -> EConstr.Unsafe.to_constr (nf_evar !evdref (EConstr.of_constr c))) fixtypes in let fiximps = List.map3 (fun ctximps cclimps (_,ctx) -> ctximps@(Impargs.lift_implicits (List.length ctx) cclimps)) fixctximps fixcclimps fixctxs in @@ -1285,9 +1294,9 @@ let do_program_recursive local p fixkind fixl ntns = let collect_evars id def typ imps = (* Generalize by the recursive prototypes *) let def = - nf_evar evd (EConstr.Unsafe.to_constr (Termops.it_mkNamedLambda_or_LetIn (EConstr.of_constr def) rec_sign)) + EConstr.Unsafe.to_constr (nf_evar evd (Termops.it_mkNamedLambda_or_LetIn (EConstr.of_constr def) rec_sign)) and typ = - nf_evar evd (EConstr.Unsafe.to_constr (Termops.it_mkNamedProd_or_LetIn (EConstr.of_constr typ) rec_sign)) + EConstr.Unsafe.to_constr (nf_evar evd (Termops.it_mkNamedProd_or_LetIn (EConstr.of_constr typ) rec_sign)) in let evm = collect_evars_of_term evd def typ in let evars, _, def, typ = diff --git a/toplevel/himsg.ml b/toplevel/himsg.ml index 3c2fe46b3..851b87903 100644 --- a/toplevel/himsg.ml +++ b/toplevel/himsg.ml @@ -86,16 +86,18 @@ let rec contract3' env a b c = function let to_unsafe_judgment j = on_judgment EConstr.Unsafe.to_constr j let j_nf_betaiotaevar sigma j = - { uj_val = Evarutil.nf_evar sigma (EConstr.Unsafe.to_constr j.uj_val); - uj_type = EConstr.Unsafe.to_constr (Reductionops.nf_betaiota sigma j.uj_type) } + { uj_val = Evarutil.nf_evar sigma j.uj_val; + uj_type = Reductionops.nf_betaiota sigma j.uj_type } let jv_nf_betaiotaevar sigma jl = - Array.map (fun j -> on_judgment EConstr.of_constr (j_nf_betaiotaevar sigma j)) jl + Array.map (fun j -> j_nf_betaiotaevar sigma j) jl (** Printers *) let pr_lconstr c = quote (pr_lconstr c) let pr_lconstr_env e s c = quote (pr_lconstr_env e s c) +let pr_leconstr c = quote (pr_leconstr c) +let pr_leconstr_env e s c = quote (pr_leconstr_env e s c) let pr_ljudge_env e s c = let v,t = pr_ljudge_env e s c in (quote v,quote t) (** A canonisation procedure for constr such that comparing there @@ -152,7 +154,8 @@ let explicit_flags = [print_implicits; print_coercions; print_no_symbol]; (** Then more! *) [print_universes; print_implicits; print_coercions; print_no_symbol] (** and more! *) ] -let pr_explicit env sigma t1 t2 = pr_explicit_aux env sigma t1 t2 explicit_flags +let pr_explicit env sigma t1 t2 = + pr_explicit_aux env sigma (EConstr.to_constr sigma t1) (EConstr.to_constr sigma t2) explicit_flags let pr_db env i = try @@ -263,9 +266,7 @@ let explain_number_branches env sigma cj expn = str "expects " ++ int expn ++ str " branches." let explain_ill_formed_branch env sigma c ci actty expty = - let actty = EConstr.Unsafe.to_constr actty in - let expty = EConstr.Unsafe.to_constr expty in - let simp t = Reduction.nf_betaiota env (Evarutil.nf_evar sigma t) in + let simp t = Reductionops.nf_betaiota sigma (Evarutil.nf_evar sigma t) in let env = make_all_name_different env in let pc = pr_leconstr_env env sigma c in let pa, pe = pr_explicit env sigma (simp actty) (simp expty) in @@ -302,13 +303,11 @@ let explain_unification_error env sigma p1 p2 = function | NotSameArgSize | NotSameHead | NoCanonicalStructure -> (* Error speaks from itself *) [] | ConversionFailed (env,t1,t2) -> - let t1 = EConstr.to_constr sigma t1 in - let t2 = EConstr.to_constr sigma t2 in - if Term.eq_constr t1 p1 && Term.eq_constr t2 p2 then [] else + if EConstr.eq_constr sigma t1 p1 && EConstr.eq_constr sigma t2 p2 then [] else let env = make_all_name_different env in let t1 = Evarutil.nf_evar sigma t1 in let t2 = Evarutil.nf_evar sigma t2 in - if not (Term.eq_constr t1 p1) || not (Term.eq_constr t2 p2) then + if not (EConstr.eq_constr sigma t1 p1) || not (EConstr.eq_constr sigma t2 p2) then let t1, t2 = pr_explicit env sigma t1 t2 in [str "cannot unify " ++ t1 ++ strbrk " and " ++ t2] else [] @@ -317,8 +316,6 @@ let explain_unification_error env sigma p1 p2 = function strbrk " refers to a metavariable - please report your example" ++ strbrk "at " ++ str Coq_config.wwwbugtracker ++ str "."] | InstanceNotSameType (evk,env,t,u) -> - let t = EConstr.to_constr sigma t in - let u = EConstr.to_constr sigma u in let t, u = pr_explicit env sigma t u in [str "unable to find a well-typed instantiation for " ++ quote (pr_existential_key sigma evk) ++ @@ -331,11 +328,13 @@ let explain_unification_error env sigma p1 p2 = function else [str "universe inconsistency"] | CannotSolveConstraint ((pb,env,t,u),e) -> + let t = EConstr.of_constr t in + let u = EConstr.of_constr u in let t = Evarutil.nf_evar sigma t in let u = Evarutil.nf_evar sigma u in let env = make_all_name_different env in - (strbrk "cannot satisfy constraint " ++ pr_lconstr_env env sigma t ++ - str " == " ++ pr_lconstr_env env sigma u) + (strbrk "cannot satisfy constraint " ++ pr_leconstr_env env sigma t ++ + str " == " ++ pr_leconstr_env env sigma u) :: aux t u e | ProblemBeyondCapabilities -> [] @@ -349,10 +348,9 @@ let explain_actual_type env sigma j t reason = let env = make_all_name_different env in let j = j_nf_betaiotaevar sigma j in let t = Reductionops.nf_betaiota sigma t in - let t = EConstr.Unsafe.to_constr t in (** Actually print *) let pe = pr_ne_context_of (str "In environment") env sigma in - let pc = pr_lconstr_env env sigma (Environ.j_val j) in + let pc = pr_leconstr_env env sigma (Environ.j_val j) in let (pt, pct) = pr_explicit env sigma t (Environ.j_type j) in let ppreason = explain_unification_error env sigma j.uj_type t reason in pe ++ @@ -363,11 +361,9 @@ let explain_actual_type env sigma j t reason = ppreason ++ str ".") let explain_cant_apply_bad_type env sigma (n,exptyp,actualtyp) rator randl = - let exptyp = EConstr.Unsafe.to_constr exptyp in let randl = jv_nf_betaiotaevar sigma randl in let exptyp = Evarutil.nf_evar sigma exptyp in let actualtyp = Reductionops.nf_betaiota sigma actualtyp in - let actualtyp = EConstr.Unsafe.to_constr actualtyp in let rator = Evarutil.j_nf_evar sigma rator in let env = make_all_name_different env in let actualtyp, exptyp = pr_explicit env sigma actualtyp exptyp in @@ -520,11 +516,9 @@ let explain_ill_formed_rec_body env sigma err names i fixenv vdefj = let explain_ill_typed_rec_body env sigma i names vdefj vargs = let vdefj = Evarutil.jv_nf_evar sigma vdefj in - let vargs = Array.map EConstr.Unsafe.to_constr vargs in - let vdefj = Array.map to_unsafe_judgment vdefj in let vargs = Array.map (Evarutil.nf_evar sigma) vargs in let env = make_all_name_different env in - let pvd = pr_lconstr_env env sigma vdefj.(i).uj_val in + let pvd = pr_leconstr_env env sigma vdefj.(i).uj_val in let pvdt, pv = pr_explicit env sigma vdefj.(i).uj_type vargs.(i) in str "The " ++ (match vdefj with [|_|] -> mt () | _ -> pr_nth (i+1) ++ spc ()) ++ @@ -584,13 +578,13 @@ let rec explain_evar_kind env sigma evk ty = function | Evar_kinds.SubEvar evk' -> let evi = Evd.find sigma evk' in let pc = match evi.evar_body with - | Evar_defined c -> pr_lconstr_env env sigma (Evarutil.nf_evar sigma c) + | Evar_defined c -> pr_leconstr_env env sigma (Evarutil.nf_evar sigma (EConstr.of_constr c)) | Evar_empty -> assert false in - let ty' = Evarutil.nf_evar sigma evi.evar_concl in + let ty' = Evarutil.nf_evar sigma (EConstr.of_constr evi.evar_concl) in pr_existential_key sigma evk ++ str " of type " ++ ty ++ str " in the partial instance " ++ pc ++ str " found for " ++ explain_evar_kind env sigma evk' - (pr_lconstr_env env sigma ty') (snd evi.evar_source) + (pr_leconstr_env env sigma ty') (snd evi.evar_source) let explain_typeclass_resolution env sigma evi k = match Typeclasses.class_of_constr sigma (EConstr.of_constr evi.evar_concl) with @@ -655,7 +649,7 @@ let explain_cannot_unify_local env sigma m n subn = let explain_refiner_cannot_generalize env sigma ty = str "Cannot find a well-typed generalisation of the goal with type: " ++ - pr_lconstr_env env sigma ty ++ str "." + pr_leconstr_env env sigma ty ++ str "." let explain_no_occurrence_found env sigma c id = let c = EConstr.to_constr sigma c in @@ -666,8 +660,8 @@ let explain_no_occurrence_found env sigma c id = | None -> str"the current goal") ++ str "." let explain_cannot_unify_binding_type env sigma m n = - let pm = pr_lconstr_env env sigma m in - let pn = pr_lconstr_env env sigma n in + let pm = pr_leconstr_env env sigma m in + let pn = pr_leconstr_env env sigma n in str "This binding has type" ++ brk(1,1) ++ pm ++ spc () ++ str "which should be unifiable with" ++ brk(1,1) ++ pn ++ str "." @@ -760,8 +754,6 @@ let explain_cannot_unify_occurrences env sigma nested ((cl2,pos2),t2) ((cl1,pos1 let ppreason = match e with | None -> mt() | Some (c1,c2,e) -> - let c1 = EConstr.to_constr sigma c1 in - let c2 = EConstr.to_constr sigma c2 in explain_unification_error env sigma c1 c2 (Some e) in str "Found incompatible occurrences of the pattern" ++ str ":" ++ @@ -841,12 +833,16 @@ let explain_pretype_error env sigma err = let actual = EConstr.Unsafe.to_constr actual in let expect = EConstr.Unsafe.to_constr expect in let env, actual, expect = contract2 env actual expect in + let actual = EConstr.of_constr actual in + let expect = EConstr.of_constr expect in explain_unexpected_type env sigma actual expect | NotProduct c -> explain_not_product env sigma c | CannotUnify (m,n,e) -> let m = EConstr.Unsafe.to_constr m in let n = EConstr.Unsafe.to_constr n in let env, m, n = contract2 env m n in + let m = EConstr.of_constr m in + let n = EConstr.of_constr n in explain_cannot_unify env sigma m n e | CannotUnifyLocal (m,n,sn) -> explain_cannot_unify_local env sigma m n sn | CannotGeneralize ty -> explain_refiner_cannot_generalize env sigma ty diff --git a/toplevel/obligations.ml b/toplevel/obligations.ml index 6da2f82ab..f4a786a73 100644 --- a/toplevel/obligations.ml +++ b/toplevel/obligations.ml @@ -818,7 +818,7 @@ let rec string_of_list sep f = function let solve_by_tac name evi t poly ctx = let id = name in - let concl = evi.evar_concl in + let concl = EConstr.of_constr evi.evar_concl in (* spiwack: the status is dropped. *) let (entry,_,ctx') = Pfedit.build_constant_by_tactic id ~goal_kind:(goal_kind poly) ctx evi.evar_hyps concl (Tacticals.New.tclCOMPLETE t) in @@ -936,7 +936,7 @@ let rec solve_obligation prg num tac = Proof_global.make_terminator (obligation_terminator prg.prg_name num guard hook auto) in let hook ctx = Lemmas.mk_hook (obligation_hook prg obl num auto ctx) in - let () = Lemmas.start_proof_univs ~sign:prg.prg_sign obl.obl_name kind evd obl.obl_type ~terminator hook in + let () = Lemmas.start_proof_univs ~sign:prg.prg_sign obl.obl_name kind evd (EConstr.of_constr obl.obl_type) ~terminator hook in let _ = Pfedit.by !default_tactic in Option.iter (fun tac -> Pfedit.set_end_tac tac) tac diff --git a/toplevel/vernacentries.ml b/toplevel/vernacentries.ml index 4376f51dc..cf3da7e02 100644 --- a/toplevel/vernacentries.ml +++ b/toplevel/vernacentries.ml @@ -59,7 +59,7 @@ let show_proof () = (* spiwack: this would probably be cooler with a bit of polishing. *) let p = Proof_global.give_me_the_proof () in let pprf = Proof.partial_proof p in - Feedback.msg_notice (Pp.prlist_with_sep Pp.fnl Printer.pr_constr pprf) + Feedback.msg_notice (Pp.prlist_with_sep Pp.fnl Printer.pr_econstr pprf) let show_node () = (* spiwack: I'm have little clue what this function used to do. I deactivated it, @@ -459,7 +459,8 @@ let start_proof_and_print k l hook = let env = Evd.evar_filtered_env evi in try let concl = Evarutil.nf_evars_universes sigma evi.Evd.evar_concl in - if Evarutil.has_undefined_evars sigma (EConstr.of_constr concl) then raise Exit; + let concl = EConstr.of_constr concl in + if Evarutil.has_undefined_evars sigma concl then raise Exit; let c, _, ctx = Pfedit.build_by_tactic env (Evd.evar_universe_context sigma) concl (Tacticals.New.tclCOMPLETE tac) @@ -877,6 +878,7 @@ let vernac_set_used_variables e = let env = Global.env () in let tys = List.map snd (Proof.initial_goals (Proof_global.give_me_the_proof ())) in + let tys = List.map EConstr.Unsafe.to_constr tys in let l = Proof_using.process_expr env e tys in let vars = Environ.named_context env in List.iter (fun id -> @@ -1901,7 +1903,7 @@ let vernac_check_guard () = try let { Evd.it=gl ; sigma=sigma } = Proof.V82.top_goal pts in Inductiveops.control_only_guard (Goal.V82.env sigma gl) - pfterm; + (EConstr.Unsafe.to_constr pfterm); (str "The condition holds up to here") with UserError(_,s) -> (str ("Condition violated: ") ++s) -- cgit v1.2.3 From 78a8d59b39dfcb07b94721fdcfd9241d404905d2 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sat, 26 Nov 2016 15:30:02 +0100 Subject: Introducing contexts parameterized by the inner term type. This allows the decoupling of the notions of context containing kernel terms and context containing tactic-level terms. --- kernel/context.ml | 69 +++++++++++--------- kernel/context.mli | 160 +++++++++++++++++++++++++--------------------- kernel/environ.ml | 2 +- kernel/indtypes.ml | 2 +- kernel/inductive.ml | 4 +- ltac/extratactics.ml4 | 2 +- pretyping/cases.ml | 2 +- pretyping/indrec.ml | 32 +++++----- pretyping/inductiveops.ml | 4 +- pretyping/typeclasses.ml | 2 +- pretyping/unification.ml | 2 +- printing/printer.ml | 4 +- printing/printmod.ml | 4 +- proofs/pfedit.ml | 2 +- tactics/class_tactics.ml | 4 +- tactics/eqschemes.ml | 58 ++++++++--------- tactics/tactics.ml | 14 ++-- toplevel/auto_ind_decl.ml | 6 +- toplevel/class.ml | 4 +- toplevel/discharge.ml | 2 +- toplevel/record.ml | 6 +- 21 files changed, 202 insertions(+), 183 deletions(-) diff --git a/kernel/context.ml b/kernel/context.ml index ae0388003..abb284f22 100644 --- a/kernel/context.ml +++ b/kernel/context.ml @@ -37,9 +37,11 @@ struct module Declaration = struct (* local declaration *) - type t = - | LocalAssum of Name.t * Constr.t (** name, type *) - | LocalDef of Name.t * Constr.t * Constr.t (** name, value, type *) + type ('constr, 'types) pt = + | LocalAssum of Name.t * 'types (** name, type *) + | LocalDef of Name.t * 'constr * 'types (** name, value, type *) + + type t = (Constr.constr, Constr.types) pt (** Return the name bound by a given declaration. *) let get_name = function @@ -87,12 +89,12 @@ struct | LocalDef (_, v, ty) -> f v && f ty (** Check whether the two given declarations are equal. *) - let equal decl1 decl2 = + let equal eq decl1 decl2 = match decl1, decl2 with | LocalAssum (n1,ty1), LocalAssum (n2, ty2) -> - Name.equal n1 n2 && Constr.equal ty1 ty2 + Name.equal n1 n2 && eq ty1 ty2 | LocalDef (n1,v1,ty1), LocalDef (n2,v2,ty2) -> - Name.equal n1 n2 && Constr.equal v1 v2 && Constr.equal ty1 ty2 + Name.equal n1 n2 && eq v1 v2 && eq ty1 ty2 | _ -> false @@ -152,6 +154,7 @@ struct (** Rel-context is represented as a list of declarations. Inner-most declarations are at the beginning of the list. Outer-most declarations are at the end of the list. *) + type ('constr, 'types) pt = ('constr, 'types) Declaration.pt list type t = Declaration.t list (** empty rel-context *) @@ -166,14 +169,14 @@ struct (** [extended_rel_list n Γ] builds an instance [args] such that [Γ,Δ ⊢ args:Γ] with n = |Δ| and with the local definitions of [Γ] skipped in [args]. Example: for [x:T,y:=c,z:U] and [n]=2, it gives [Rel 5, Rel 3]. *) - let nhyps = + let nhyps ctx = let open Declaration in let rec nhyps acc = function | [] -> acc | LocalAssum _ :: hyps -> nhyps (succ acc) hyps | LocalDef _ :: hyps -> nhyps acc hyps in - nhyps 0 + nhyps 0 ctx (** Return a declaration designated by a given de Bruijn index. @raise Not_found if the designated de Bruijn index is not present in the designated rel-context. *) @@ -184,7 +187,7 @@ struct | _, [] -> raise Not_found (** Check whether given two rel-contexts are equal. *) - let equal = List.equal Declaration.equal + let equal eq l = List.equal (fun c -> Declaration.equal eq c) l (** Map all terms in a given rel-context. *) let map f = List.smartmap (Declaration.map_constr f) @@ -202,26 +205,26 @@ struct (** Map a given rel-context to a list where each {e local assumption} is mapped to [true] and each {e local definition} is mapped to [false]. *) - let to_tags = + let to_tags l = let rec aux l = function | [] -> l | Declaration.LocalDef _ :: ctx -> aux (true::l) ctx | Declaration.LocalAssum _ :: ctx -> aux (false::l) ctx - in aux [] + in aux [] l (** [extended_list n Γ] builds an instance [args] such that [Γ,Δ ⊢ args:Γ] with n = |Δ| and with the {e local definitions} of [Γ] skipped in [args]. Example: for [x:T, y:=c, z:U] and [n]=2, it gives [Rel 5, Rel 3]. *) - let to_extended_list n = + let to_extended_list mk n l = let rec reln l p = function - | Declaration.LocalAssum _ :: hyps -> reln (Constr.mkRel (n+p) :: l) (p+1) hyps + | Declaration.LocalAssum _ :: hyps -> reln (mk (n+p) :: l) (p+1) hyps | Declaration.LocalDef _ :: hyps -> reln l (p+1) hyps | [] -> l in - reln [] 1 + reln [] 1 l (** [extended_vect n Γ] does the same, returning instead an array. *) - let to_extended_vect n hyps = Array.of_list (to_extended_list n hyps) + let to_extended_vect mk n hyps = Array.of_list (to_extended_list mk n hyps) end (** This module represents contexts that can capture non-anonymous variables. @@ -232,9 +235,11 @@ struct module Declaration = struct (** local declaration *) - type t = - | LocalAssum of Id.t * Constr.t (** identifier, type *) - | LocalDef of Id.t * Constr.t * Constr.t (** identifier, value, type *) + type ('constr, 'types) pt = + | LocalAssum of Id.t * 'types (** identifier, type *) + | LocalDef of Id.t * 'constr * 'types (** identifier, value, type *) + + type t = (Constr.constr, Constr.types) pt (** Return the identifier bound by a given declaration. *) let get_id = function @@ -282,12 +287,12 @@ struct | LocalDef (_, v, ty) -> f v && f ty (** Check whether the two given declarations are equal. *) - let equal decl1 decl2 = + let equal eq decl1 decl2 = match decl1, decl2 with | LocalAssum (id1, ty1), LocalAssum (id2, ty2) -> - Id.equal id1 id2 && Constr.equal ty1 ty2 + Id.equal id1 id2 && eq ty1 ty2 | LocalDef (id1, v1, ty1), LocalDef (id2, v2, ty2) -> - Id.equal id1 id2 && Constr.equal v1 v2 && Constr.equal ty1 ty2 + Id.equal id1 id2 && eq v1 v2 && eq ty1 ty2 | _ -> false @@ -362,6 +367,7 @@ struct (** Named-context is represented as a list of declarations. Inner-most declarations are at the beginning of the list. Outer-most declarations are at the end of the list. *) + type ('constr, 'types) pt = ('constr, 'types) Declaration.pt list type t = Declaration.t list (** empty named-context *) @@ -380,7 +386,7 @@ struct | [] -> raise Not_found (** Check whether given two named-contexts are equal. *) - let equal = List.equal Declaration.equal + let equal eq l = List.equal (fun c -> Declaration.equal eq c) l (** Map all terms in a given named-context. *) let map f = List.smartmap (Declaration.map_constr f) @@ -397,28 +403,30 @@ struct let fold_outside f l ~init = List.fold_right f l init (** Return the set of all identifiers bound in a given named-context. *) - let to_vars = - List.fold_left (fun accu decl -> Id.Set.add (Declaration.get_id decl) accu) Id.Set.empty + let to_vars l = + List.fold_left (fun accu decl -> Id.Set.add (Declaration.get_id decl) accu) Id.Set.empty l (** [instance_from_named_context Ω] builds an instance [args] such that [Ω ⊢ args:Ω] where [Ω] is a named context and with the local definitions of [Ω] skipped. Example: for [id1:T,id2:=c,id3:U], it gives [Var id1, Var id3]. All [idj] are supposed distinct. *) - let to_instance = + let to_instance mk l = let filter = function - | Declaration.LocalAssum (id, _) -> Some (Constr.mkVar id) + | Declaration.LocalAssum (id, _) -> Some (mk id) | _ -> None in - List.map_filter filter + List.map_filter filter l end module Compacted = struct module Declaration = struct - type t = - | LocalAssum of Id.t list * Constr.t - | LocalDef of Id.t list * Constr.t * Constr.t + type ('constr, 'types) pt = + | LocalAssum of Id.t list * 'types + | LocalDef of Id.t list * 'constr * 'types + + type t = (Constr.constr, Constr.types) pt let map_constr f = function | LocalAssum (ids, ty) as decl -> @@ -442,6 +450,7 @@ module Compacted = List.map (fun id -> Named.Declaration.LocalDef (id,v,t)) ids end + type ('constr, 'types) pt = ('constr, 'types) Declaration.pt list type t = Declaration.t list let fold f l ~init = List.fold_right f l init diff --git a/kernel/context.mli b/kernel/context.mli index 955e214cb..0c666a25d 100644 --- a/kernel/context.mli +++ b/kernel/context.mli @@ -29,110 +29,115 @@ sig module Declaration : sig (* local declaration *) - type t = LocalAssum of Name.t * Constr.t (** name, type *) - | LocalDef of Name.t * Constr.t * Constr.t (** name, value, type *) + type ('constr, 'types) pt = + | LocalAssum of Name.t * 'types (** name, type *) + | LocalDef of Name.t * 'constr * 'types (** name, value, type *) + + type t = (Constr.constr, Constr.types) pt (** Return the name bound by a given declaration. *) - val get_name : t -> Name.t + val get_name : ('c, 't) pt -> Name.t (** Return [Some value] for local-declarations and [None] for local-assumptions. *) - val get_value : t -> Constr.t option + val get_value : ('c, 't) pt -> 'c option (** Return the type of the name bound by a given declaration. *) - val get_type : t -> Constr.t + val get_type : ('c, 't) pt -> 't (** Set the name that is bound by a given declaration. *) - val set_name : Name.t -> t -> t + val set_name : Name.t -> ('c, 't) pt -> ('c, 't) pt (** Set the type of the bound variable in a given declaration. *) - val set_type : Constr.t -> t -> t + val set_type : 't -> ('c, 't) pt -> ('c, 't) pt (** Return [true] iff a given declaration is a local assumption. *) - val is_local_assum : t -> bool + val is_local_assum : ('c, 't) pt -> bool (** Return [true] iff a given declaration is a local definition. *) - val is_local_def : t -> bool + val is_local_def : ('c, 't) pt -> bool (** Check whether any term in a given declaration satisfies a given predicate. *) - val exists : (Constr.t -> bool) -> t -> bool + val exists : ('c -> bool) -> ('c, 'c) pt -> bool (** Check whether all terms in a given declaration satisfy a given predicate. *) - val for_all : (Constr.t -> bool) -> t -> bool + val for_all : ('c -> bool) -> ('c, 'c) pt -> bool (** Check whether the two given declarations are equal. *) - val equal : t -> t -> bool + val equal : ('c -> 'c -> bool) -> ('c, 'c) pt -> ('c, 'c) pt -> bool (** Map the name bound by a given declaration. *) - val map_name : (Name.t -> Name.t) -> t -> t + val map_name : (Name.t -> Name.t) -> ('c, 't) pt -> ('c, 't) pt (** For local assumptions, this function returns the original local assumptions. For local definitions, this function maps the value in the local definition. *) - val map_value : (Constr.t -> Constr.t) -> t -> t + val map_value : ('c -> 'c) -> ('c, 't) pt -> ('c, 't) pt (** Map the type of the name bound by a given declaration. *) - val map_type : (Constr.t -> Constr.t) -> t -> t + val map_type : ('t -> 't) -> ('c, 't) pt -> ('c, 't) pt (** Map all terms in a given declaration. *) - val map_constr : (Constr.t -> Constr.t) -> t -> t + val map_constr : ('c -> 'c) -> ('c, 'c) pt -> ('c, 'c) pt (** Perform a given action on all terms in a given declaration. *) - val iter_constr : (Constr.t -> unit) -> t -> unit + val iter_constr : ('c -> unit) -> ('c, 'c) pt -> unit (** Reduce all terms in a given declaration to a single value. *) - val fold_constr : (Constr.t -> 'a -> 'a) -> t -> 'a -> 'a + val fold_constr : ('c -> 'a -> 'a) -> ('c, 'c) pt -> 'a -> 'a - val to_tuple : t -> Name.t * Constr.t option * Constr.t + val to_tuple : ('c, 't) pt -> Name.t * 'c option * 't end (** Rel-context is represented as a list of declarations. Inner-most declarations are at the beginning of the list. Outer-most declarations are at the end of the list. *) + type ('constr, 'types) pt = ('constr, 'types) Declaration.pt list type t = Declaration.t list (** empty rel-context *) - val empty : t + val empty : ('c, 't) pt (** Return a new rel-context enriched by with a given inner-most declaration. *) - val add : Declaration.t -> t -> t + val add : ('c, 't) Declaration.pt -> ('c, 't) pt -> ('c, 't) pt (** Return the number of {e local declarations} in a given context. *) - val length : t -> int + val length : ('c, 't) pt -> int (** Check whether given two rel-contexts are equal. *) - val equal : t -> t -> bool + val equal : ('c -> 'c -> bool) -> ('c, 'c) pt -> ('c, 'c) pt -> bool (** Return the number of {e local assumptions} in a given rel-context. *) - val nhyps : t -> int + val nhyps : ('c, 't) pt -> int (** Return a declaration designated by a given de Bruijn index. @raise Not_found if the designated de Bruijn index outside the range. *) - val lookup : int -> t -> Declaration.t + val lookup : int -> ('c, 't) pt -> ('c, 't) Declaration.pt (** Map all terms in a given rel-context. *) - val map : (Constr.t -> Constr.t) -> t -> t + val map : ('c -> 'c) -> ('c, 'c) pt -> ('c, 'c) pt (** Perform a given action on every declaration in a given rel-context. *) - val iter : (Constr.t -> unit) -> t -> unit + val iter : ('c -> unit) -> ('c, 'c) pt -> unit (** Reduce all terms in a given rel-context to a single value. Innermost declarations are processed first. *) - val fold_inside : ('a -> Declaration.t -> 'a) -> init:'a -> t -> 'a + val fold_inside : ('a -> ('c, 't) Declaration.pt -> 'a) -> init:'a -> ('c, 't) pt -> 'a (** Reduce all terms in a given rel-context to a single value. Outermost declarations are processed first. *) - val fold_outside : (Declaration.t -> 'a -> 'a) -> t -> init:'a -> 'a + val fold_outside : (('c, 't) Declaration.pt -> 'a -> 'a) -> ('c, 't) pt -> init:'a -> 'a (** Map a given rel-context to a list where each {e local assumption} is mapped to [true] and each {e local definition} is mapped to [false]. *) - val to_tags : t -> bool list + val to_tags : ('c, 't) pt -> bool list - (** [extended_list n Γ] builds an instance [args] such that [Γ,Δ ⊢ args:Γ] + (** [extended_list mk n Γ] builds an instance [args] such that [Γ,Δ ⊢ args:Γ] with n = |Δ| and with the {e local definitions} of [Γ] skipped in - [args]. Example: for [x:T, y:=c, z:U] and [n]=2, it gives [Rel 5, Rel 3]. *) - val to_extended_list : int -> t -> Constr.t list + [args] where [mk] is used to build the corresponding variables. + Example: for [x:T, y:=c, z:U] and [n]=2, it gives [mk 5, mk 3]. *) + val to_extended_list : (int -> 'r) -> int -> ('c, 't) pt -> 'r list (** [extended_vect n Γ] does the same, returning instead an array. *) - val to_extended_vect : int -> t -> Constr.t array + val to_extended_vect : (int -> 'r) -> int -> ('c, 't) pt -> 'r array end (** This module represents contexts that can capture non-anonymous variables. @@ -142,129 +147,136 @@ sig (** Representation of {e local declarations}. *) module Declaration : sig - type t = LocalAssum of Id.t * Constr.t (** identifier, type *) - | LocalDef of Id.t * Constr.t * Constr.t (** identifier, value, type *) + type ('constr, 'types) pt = + | LocalAssum of Id.t * 'types (** identifier, type *) + | LocalDef of Id.t * 'constr * 'types (** identifier, value, type *) + + type t = (Constr.constr, Constr.types) pt (** Return the identifier bound by a given declaration. *) - val get_id : t -> Id.t + val get_id : ('c, 't) pt -> Id.t (** Return [Some value] for local-declarations and [None] for local-assumptions. *) - val get_value : t -> Constr.t option + val get_value : ('c, 't) pt -> 'c option (** Return the type of the name bound by a given declaration. *) - val get_type : t -> Constr.t + val get_type : ('c, 't) pt -> 't (** Set the identifier that is bound by a given declaration. *) - val set_id : Id.t -> t -> t + val set_id : Id.t -> ('c, 't) pt -> ('c, 't) pt (** Set the type of the bound variable in a given declaration. *) - val set_type : Constr.t -> t -> t + val set_type : 't -> ('c, 't) pt -> ('c, 't) pt (** Return [true] iff a given declaration is a local assumption. *) - val is_local_assum : t -> bool + val is_local_assum : ('c, 't) pt -> bool (** Return [true] iff a given declaration is a local definition. *) - val is_local_def : t -> bool + val is_local_def : ('c, 't) pt -> bool (** Check whether any term in a given declaration satisfies a given predicate. *) - val exists : (Constr.t -> bool) -> t -> bool + val exists : ('c -> bool) -> ('c, 'c) pt -> bool (** Check whether all terms in a given declaration satisfy a given predicate. *) - val for_all : (Constr.t -> bool) -> t -> bool + val for_all : ('c -> bool) -> ('c, 'c) pt -> bool (** Check whether the two given declarations are equal. *) - val equal : t -> t -> bool + val equal : ('c -> 'c -> bool) -> ('c, 'c) pt -> ('c, 'c) pt -> bool (** Map the identifier bound by a given declaration. *) - val map_id : (Id.t -> Id.t) -> t -> t + val map_id : (Id.t -> Id.t) -> ('c, 't) pt -> ('c, 't) pt (** For local assumptions, this function returns the original local assumptions. For local definitions, this function maps the value in the local definition. *) - val map_value : (Constr.t -> Constr.t) -> t -> t + val map_value : ('c -> 'c) -> ('c, 't) pt -> ('c, 't) pt (** Map the type of the name bound by a given declaration. *) - val map_type : (Constr.t -> Constr.t) -> t -> t + val map_type : ('t -> 't) -> ('c, 't) pt -> ('c, 't) pt (** Map all terms in a given declaration. *) - val map_constr : (Constr.t -> Constr.t) -> t -> t + val map_constr : ('c -> 'c) -> ('c, 'c) pt -> ('c, 'c) pt (** Perform a given action on all terms in a given declaration. *) - val iter_constr : (Constr.t -> unit) -> t -> unit + val iter_constr : ('c -> unit) -> ('c, 'c) pt -> unit (** Reduce all terms in a given declaration to a single value. *) - val fold_constr : (Constr.t -> 'a -> 'a) -> t -> 'a -> 'a + val fold_constr : ('c -> 'a -> 'a) -> ('c, 'c) pt -> 'a -> 'a - val to_tuple : t -> Id.t * Constr.t option * Constr.t - val of_tuple : Id.t * Constr.t option * Constr.t -> t + val to_tuple : ('c, 't) pt -> Id.t * 'c option * 't + val of_tuple : Id.t * 'c option * 't -> ('c, 't) pt (** Convert [Rel.Declaration.t] value to the corresponding [Named.Declaration.t] value. The function provided as the first parameter determines how to translate "names" to "ids". *) - val of_rel_decl : (Name.t -> Id.t) -> Rel.Declaration.t -> t + val of_rel_decl : (Name.t -> Id.t) -> ('c, 't) Rel.Declaration.pt -> ('c, 't) pt (** Convert [Named.Declaration.t] value to the corresponding [Rel.Declaration.t] value. *) (* TODO: Move this function to [Rel.Declaration] module and rename it to [of_named]. *) - val to_rel_decl : t -> Rel.Declaration.t + val to_rel_decl : ('c, 't) pt -> ('c, 't) Rel.Declaration.pt end (** Rel-context is represented as a list of declarations. Inner-most declarations are at the beginning of the list. Outer-most declarations are at the end of the list. *) + type ('constr, 'types) pt = ('constr, 'types) Declaration.pt list type t = Declaration.t list (** empty named-context *) - val empty : t + val empty : ('c, 't) pt (** Return a new rel-context enriched by with a given inner-most declaration. *) - val add : Declaration.t -> t -> t + val add : ('c, 't) Declaration.pt -> ('c, 't) pt -> ('c, 't) pt (** Return the number of {e local declarations} in a given named-context. *) - val length : t -> int + val length : ('c, 't) pt -> int (** Return a declaration designated by an identifier of the variable bound in that declaration. @raise Not_found if the designated identifier is not bound in a given named-context. *) - val lookup : Id.t -> t -> Declaration.t + val lookup : Id.t -> ('c, 't) pt -> ('c, 't) Declaration.pt (** Check whether given two rel-contexts are equal. *) - val equal : t -> t -> bool + val equal : ('c -> 'c -> bool) -> ('c, 'c) pt -> ('c, 'c) pt -> bool (** Map all terms in a given named-context. *) - val map : (Constr.t -> Constr.t) -> t -> t + val map : ('c -> 'c) -> ('c, 'c) pt -> ('c, 'c) pt (** Perform a given action on every declaration in a given named-context. *) - val iter : (Constr.t -> unit) -> t -> unit + val iter : ('c -> unit) -> ('c, 'c) pt -> unit (** Reduce all terms in a given named-context to a single value. Innermost declarations are processed first. *) - val fold_inside : ('a -> Declaration.t -> 'a) -> init:'a -> t -> 'a + val fold_inside : ('a -> ('c, 't) Declaration.pt -> 'a) -> init:'a -> ('c, 't) pt -> 'a (** Reduce all terms in a given named-context to a single value. Outermost declarations are processed first. *) - val fold_outside : (Declaration.t -> 'a -> 'a) -> t -> init:'a -> 'a + val fold_outside : (('c, 't) Declaration.pt -> 'a -> 'a) -> ('c, 't) pt -> init:'a -> 'a (** Return the set of all identifiers bound in a given named-context. *) - val to_vars : t -> Id.Set.t + val to_vars : ('c, 't) pt -> Id.Set.t (** [instance_from_named_context Ω] builds an instance [args] such that [Ω ⊢ args:Ω] where [Ω] is a named context and with the local definitions of [Ω] skipped. Example: for [id1:T,id2:=c,id3:U], it gives [Var id1, Var id3]. All [idj] are supposed distinct. *) - val to_instance : t -> Constr.t list + val to_instance : (Id.t -> 'r) -> ('c, 't) pt -> 'r list end module Compacted : sig module Declaration : sig - type t = - | LocalAssum of Id.t list * Constr.t - | LocalDef of Id.t list * Constr.t * Constr.t + type ('constr, 'types) pt = + | LocalAssum of Id.t list * 'types + | LocalDef of Id.t list * 'constr * 'types + + type t = (Constr.constr, Constr.types) pt - val map_constr : (Constr.t -> Constr.t) -> t -> t - val of_named_decl : Named.Declaration.t -> t - val to_named_context : t -> Named.t + val map_constr : ('c -> 'c) -> ('c, 'c) pt -> ('c, 'c) pt + val of_named_decl : ('c, 't) Named.Declaration.pt -> ('c, 't) pt + val to_named_context : ('c, 't) pt -> ('c, 't) Named.pt end + type ('constr, 'types) pt = ('constr, 'types) Declaration.pt list type t = Declaration.t list - val fold : (Declaration.t -> 'a -> 'a) -> t -> init:'a -> 'a + val fold : (('c, 't) Declaration.pt -> 'a -> 'a) -> ('c, 't) pt -> init:'a -> 'a end diff --git a/kernel/environ.ml b/kernel/environ.ml index 5688813ad..9986f439a 100644 --- a/kernel/environ.ml +++ b/kernel/environ.ml @@ -120,7 +120,7 @@ let lookup_named = lookup_named let lookup_named_val id ctxt = fst (Id.Map.find id ctxt.env_named_map) let eq_named_context_val c1 c2 = - c1 == c2 || Context.Named.equal (named_context_of_val c1) (named_context_of_val c2) + c1 == c2 || Context.Named.equal Constr.equal (named_context_of_val c1) (named_context_of_val c2) (* A local const is evaluable if it is defined *) diff --git a/kernel/indtypes.ml b/kernel/indtypes.ml index de97268b3..2ff419338 100644 --- a/kernel/indtypes.ml +++ b/kernel/indtypes.ml @@ -734,7 +734,7 @@ let compute_projections ((kn, _ as ind), u as indu) n x nparamargs params matching with a parameter context. *) let indty, paramsletsubst = (* [ty] = [Ind inst] is typed in context [params] *) - let inst = Context.Rel.to_extended_vect 0 paramslet in + let inst = Context.Rel.to_extended_vect mkRel 0 paramslet in let ty = mkApp (mkIndU indu, inst) in (* [Ind inst] is typed in context [params-wo-let] *) let inst' = rel_list 0 nparamargs in diff --git a/kernel/inductive.ml b/kernel/inductive.ml index 3c4c2796e..d8d365c34 100644 --- a/kernel/inductive.ml +++ b/kernel/inductive.ml @@ -297,7 +297,7 @@ let build_dependent_inductive ind (_,mip) params = applist (mkIndU ind, List.map (lift mip.mind_nrealdecls) params - @ Context.Rel.to_extended_list 0 realargs) + @ Context.Rel.to_extended_list mkRel 0 realargs) (* This exception is local *) exception LocalArity of (sorts_family * sorts_family * arity_error) option @@ -355,7 +355,7 @@ let build_branches_type (ind,u) (_,mip as specif) params p = let (lparams,vargs) = List.chop (inductive_params specif) allargs in let cargs = let cstr = ith_constructor_of_inductive ind (i+1) in - let dep_cstr = applist (mkConstructU (cstr,u),lparams@(Context.Rel.to_extended_list 0 cstrsign)) in + let dep_cstr = applist (mkConstructU (cstr,u),lparams@(Context.Rel.to_extended_list mkRel 0 cstrsign)) in vargs @ [dep_cstr] in let base = lambda_appvect_assum (mip.mind_nrealdecls+1) (lift nargs p) (Array.of_list cargs) in it_mkProd_or_LetIn base cstrsign in diff --git a/ltac/extratactics.ml4 b/ltac/extratactics.ml4 index 519633bbe..34159d945 100644 --- a/ltac/extratactics.ml4 +++ b/ltac/extratactics.ml4 @@ -310,7 +310,7 @@ let project_hint pri l2r r = let p = if l2r then build_coq_iff_left_proj () else build_coq_iff_right_proj () in let p = EConstr.of_constr p in - let c = Reductionops.whd_beta sigma (mkApp (c, Array.map EConstr.of_constr (Context.Rel.to_extended_vect 0 sign))) in + let c = Reductionops.whd_beta sigma (mkApp (c, Context.Rel.to_extended_vect mkRel 0 sign)) in let c = it_mkLambda_or_LetIn (mkApp (p,[|mkArrow a (lift 1 b);mkArrow b (lift 1 a);c|])) sign in let id = diff --git a/pretyping/cases.ml b/pretyping/cases.ml index 0e4c6619b..3b5662a24 100644 --- a/pretyping/cases.ml +++ b/pretyping/cases.ml @@ -943,7 +943,7 @@ let abstract_predicate env sigma indf cur realargs (names,na) tms ccl = let tms = List.fold_right2 (fun par arg tomatch -> match EConstr.kind sigma par with | Rel i -> relocate_index_tomatch sigma (i+n) (destRel sigma arg) tomatch - | _ -> tomatch) (realargs@[cur]) (List.map EConstr.of_constr (Context.Rel.to_extended_list 0 sign)) + | _ -> tomatch) (realargs@[cur]) (Context.Rel.to_extended_list EConstr.mkRel 0 sign) (lift_tomatch_stack n tms) in (* Pred is already dependent in the current term to match (if *) (* (na<>Anonymous) and its realargs; we just need to adjust it to *) diff --git a/pretyping/indrec.ml b/pretyping/indrec.ml index 1adeb4db2..431d1ff16 100644 --- a/pretyping/indrec.ml +++ b/pretyping/indrec.ml @@ -63,7 +63,7 @@ let check_privacy_block mib = let mis_make_case_com dep env sigma (ind, u as pind) (mib,mip as specif) kind = let lnamespar = Vars.subst_instance_context u mib.mind_params_ctxt in - let indf = make_ind_family(pind, Context.Rel.to_extended_list 0 lnamespar) in + let indf = make_ind_family(pind, Context.Rel.to_extended_list mkRel 0 lnamespar) in let constrs = get_constructors env indf in let projs = get_projections env indf in @@ -93,8 +93,8 @@ let mis_make_case_com dep env sigma (ind, u as pind) (mib,mip as specif) kind = let pbody = appvect (mkRel (ndepar + nbprod), - if dep then Context.Rel.to_extended_vect 0 deparsign - else Context.Rel.to_extended_vect 1 arsign) in + if dep then Context.Rel.to_extended_vect mkRel 0 deparsign + else Context.Rel.to_extended_vect mkRel 1 arsign) in let p = it_mkLambda_or_LetIn_name env' ((if dep then mkLambda_name env' else mkLambda) @@ -168,7 +168,7 @@ let type_rec_branch is_rec dep env sigma (vargs,depPvect,decP) tyi cs recargs = let base = applist (lift i pk,realargs) in if depK then Reduction.beta_appvect - base [|applist (mkRel (i+1), Context.Rel.to_extended_list 0 sign)|] + base [|applist (mkRel (i+1), Context.Rel.to_extended_list mkRel 0 sign)|] else base | _ -> @@ -244,7 +244,7 @@ let make_rec_branch_arg env sigma (nparrec,fvect,decF) f cstr recargs = mkLetIn (n,b,t,prec (push_rel d env) (i+1) (d::hyps) c) | Ind _ -> let realargs = List.skipn nparrec largs - and arg = appvect (mkRel (i+1), Context.Rel.to_extended_vect 0 hyps) in + and arg = appvect (mkRel (i+1), Context.Rel.to_extended_vect mkRel 0 hyps) in applist(lift i fk,realargs@[arg]) | _ -> let t' = whd_all env sigma (EConstr.of_constr p) in @@ -323,7 +323,7 @@ let mis_make_indrec env sigma listdepkind mib u = (* arity in the context of the fixpoint, i.e. P1..P_nrec f1..f_nbconstruct *) - let args = Context.Rel.to_extended_list (nrec+nbconstruct) lnamesparrec in + let args = Context.Rel.to_extended_list mkRel (nrec+nbconstruct) lnamesparrec in let indf = make_ind_family((indi,u),args) in let arsign,_ = get_arity env indf in @@ -337,15 +337,15 @@ let mis_make_indrec env sigma listdepkind mib u = (* constructors in context of the Cases expr, i.e. P1..P_nrec f1..f_nbconstruct F_1..F_nrec a_1..a_nar x:I *) - let args' = Context.Rel.to_extended_list (dect+nrec) lnamesparrec in - let args'' = Context.Rel.to_extended_list ndepar lnonparrec in + let args' = Context.Rel.to_extended_list mkRel (dect+nrec) lnamesparrec in + let args'' = Context.Rel.to_extended_list mkRel ndepar lnonparrec in let indf' = make_ind_family((indi,u),args'@args'') in let branches = let constrs = get_constructors env indf' in let fi = Termops.rel_vect (dect-i-nctyi) nctyi in let vecfi = Array.map - (fun f -> appvect (f, Context.Rel.to_extended_vect ndepar lnonparrec)) + (fun f -> appvect (f, Context.Rel.to_extended_vect mkRel ndepar lnonparrec)) fi in Array.map3 @@ -366,9 +366,9 @@ let mis_make_indrec env sigma listdepkind mib u = let deparsign' = LocalAssum (Anonymous,depind')::arsign' in let pargs = - let nrpar = Context.Rel.to_extended_list (2*ndepar) lnonparrec - and nrar = if dep then Context.Rel.to_extended_list 0 deparsign' - else Context.Rel.to_extended_list 1 arsign' + let nrpar = Context.Rel.to_extended_list mkRel (2*ndepar) lnonparrec + and nrar = if dep then Context.Rel.to_extended_list mkRel 0 deparsign' + else Context.Rel.to_extended_list mkRel 1 arsign' in nrpar@nrar in @@ -396,8 +396,8 @@ let mis_make_indrec env sigma listdepkind mib u = let typtyi = let concl = - let pargs = if dep then Context.Rel.to_extended_vect 0 deparsign - else Context.Rel.to_extended_vect 1 arsign + let pargs = if dep then Context.Rel.to_extended_vect mkRel 0 deparsign + else Context.Rel.to_extended_vect mkRel 1 arsign in appvect (mkRel (nbconstruct+ndepar+nonrecpar+j),pargs) in it_mkProd_or_LetIn_name env concl @@ -424,7 +424,7 @@ let mis_make_indrec env sigma listdepkind mib u = else let recarg = (dest_subterms recargsvec.(tyi)).(j) in let recarg = recargpar@recarg in - let vargs = Context.Rel.to_extended_list (nrec+i+j) lnamesparrec in + let vargs = Context.Rel.to_extended_list mkRel (nrec+i+j) lnamesparrec in let cs = get_constructor ((indi,u),mibi,mipi,vargs) (j+1) in let p_0 = type_rec_branch @@ -438,7 +438,7 @@ let mis_make_indrec env sigma listdepkind mib u = in let rec put_arity env i = function | ((indi,u),_,_,dep,kinds)::rest -> - let indf = make_ind_family ((indi,u), Context.Rel.to_extended_list i lnamesparrec) in + let indf = make_ind_family ((indi,u), Context.Rel.to_extended_list mkRel i lnamesparrec) in let s = Evarutil.evd_comb1 (Evd.fresh_sort_in_family ~rigid:Evd.univ_flexible_alg env) evdref kinds diff --git a/pretyping/inductiveops.ml b/pretyping/inductiveops.ml index 1dcd6399e..c00ceb02e 100644 --- a/pretyping/inductiveops.ml +++ b/pretyping/inductiveops.ml @@ -411,14 +411,14 @@ let build_dependent_constructor cs = applist (mkConstructU cs.cs_cstr, (List.map (lift cs.cs_nargs) cs.cs_params) - @(Context.Rel.to_extended_list 0 cs.cs_args)) + @(Context.Rel.to_extended_list mkRel 0 cs.cs_args)) let build_dependent_inductive env ((ind, params) as indf) = let arsign,_ = get_arity env indf in let nrealargs = List.length arsign in applist (mkIndU ind, - (List.map (lift nrealargs) params)@(Context.Rel.to_extended_list 0 arsign)) + (List.map (lift nrealargs) params)@(Context.Rel.to_extended_list mkRel 0 arsign)) (* builds the arity of an elimination predicate in sort [s] *) diff --git a/pretyping/typeclasses.ml b/pretyping/typeclasses.ml index 9da7005e0..50ae66eb0 100644 --- a/pretyping/typeclasses.ml +++ b/pretyping/typeclasses.ml @@ -289,7 +289,7 @@ let build_subclasses ~check env sigma glob pri = | None -> [] | Some (rels, ((tc,u), args)) -> let instapp = - Reductionops.whd_beta sigma (EConstr.of_constr (appvectc c (Context.Rel.to_extended_vect 0 rels))) + Reductionops.whd_beta sigma (EConstr.of_constr (appvectc c (Context.Rel.to_extended_vect mkRel 0 rels))) in let instapp = EConstr.Unsafe.to_constr instapp in let projargs = Array.of_list (args @ [instapp]) in diff --git a/pretyping/unification.ml b/pretyping/unification.ml index 1dc3ccdc0..04cc4253e 100644 --- a/pretyping/unification.ml +++ b/pretyping/unification.ml @@ -1634,7 +1634,7 @@ let make_abstraction_core name (test,out) env sigma c ty occs check_occs concl = | AllOccurrences, InHyp as occ -> let occ = if likefirst then LikeFirst else AtOccs occ in let newdecl = replace_term_occ_decl_modulo sigma occ test mkvarid d in - if Context.Named.Declaration.equal d newdecl + if Context.Named.Declaration.equal Constr.equal d newdecl && not (indirectly_dependent sigma c d depdecls) then if check_occs && not (in_every_hyp occs) diff --git a/printing/printer.ml b/printing/printer.ml index 63aeec82c..d9060c172 100644 --- a/printing/printer.ml +++ b/printing/printer.ml @@ -499,8 +499,8 @@ let print_evar_constraints gl sigma = | Some g -> let env = Goal.V82.env sigma g in fun e' -> begin - if Context.Named.equal (named_context env) (named_context e') then - if Context.Rel.equal (rel_context env) (rel_context e') then mt () + if Context.Named.equal Constr.equal (named_context env) (named_context e') then + if Context.Rel.equal Constr.equal (rel_context env) (rel_context e') then mt () else pr_rel_context_of e' sigma ++ str " |-" ++ spc () else pr_context_of e' sigma ++ str " |-" ++ spc () end diff --git a/printing/printmod.ml b/printing/printmod.ml index dfa66d437..e8ea0a99a 100644 --- a/printing/printmod.ml +++ b/printing/printmod.ml @@ -88,7 +88,7 @@ let print_one_inductive env sigma mib ((_,i) as ind) = else Univ.Instance.empty in let mip = mib.mind_packets.(i) in let params = Inductive.inductive_paramdecls (mib,u) in - let args = Context.Rel.to_extended_list 0 params in + let args = Context.Rel.to_extended_list mkRel 0 params in let arity = hnf_prod_applist env (build_ind_type env ((mib,mip),u)) args in let cstrtypes = Inductive.type_of_constructors (ind,u) (mib,mip) in let cstrtypes = Array.map (fun c -> hnf_prod_applist env c args) cstrtypes in @@ -142,7 +142,7 @@ let print_record env mind mib = in let mip = mib.mind_packets.(0) in let params = Inductive.inductive_paramdecls (mib,u) in - let args = Context.Rel.to_extended_list 0 params in + let args = Context.Rel.to_extended_list mkRel 0 params in let arity = hnf_prod_applist env (build_ind_type env ((mib,mip),u)) args in let cstrtypes = Inductive.type_of_constructors ((mind,0),u) (mib,mip) in let cstrtype = hnf_prod_applist env cstrtypes.(0) args in diff --git a/proofs/pfedit.ml b/proofs/pfedit.ml index 142fd9a89..09b924c7e 100644 --- a/proofs/pfedit.ml +++ b/proofs/pfedit.ml @@ -228,7 +228,7 @@ let solve_by_implicit_tactic env sigma evk = match (!implicit_tactic, snd (evar_source evk sigma)) with | Some tac, (Evar_kinds.ImplicitArg _ | Evar_kinds.QuestionMark _) when - Context.Named.equal (Environ.named_context_of_val evi.evar_hyps) + Context.Named.equal Constr.equal (Environ.named_context_of_val evi.evar_hyps) (Environ.named_context env) -> let tac = Proofview.tclTHEN tac (Proofview.tclEXTEND [] (Proofview.tclZERO (CErrors.UserError (None,Pp.str"Proof is not complete."))) []) in (try diff --git a/tactics/class_tactics.ml b/tactics/class_tactics.ml index a4b6cb53b..ef67d28f9 100644 --- a/tactics/class_tactics.ml +++ b/tactics/class_tactics.ml @@ -939,7 +939,7 @@ module Search = struct let cwd = Lib.cwd () in if DirPath.equal cwd dir && (onlyc == only_classes) && - Context.Named.equal sign sign' && + Context.Named.equal Constr.equal sign sign' && Hint_db.transparent_state cached_hints == st then cached_hints else @@ -1034,7 +1034,7 @@ module Search = struct (pr_depth (succ j :: i :: info.search_depth) ++ str" : " ++ pr_ev s' (Proofview.Goal.goal gl')); let hints' = - if b && not (Context.Named.equal (Goal.hyps gl') (Goal.hyps gl)) + if b && not (Context.Named.equal Constr.equal (Goal.hyps gl') (Goal.hyps gl)) then let st = Hint_db.transparent_state info.search_hints in make_autogoal_hints info.search_only_classes ~st gl' diff --git a/tactics/eqschemes.ml b/tactics/eqschemes.ml index 855273d3b..188e215a5 100644 --- a/tactics/eqschemes.ml +++ b/tactics/eqschemes.ml @@ -73,8 +73,8 @@ let build_dependent_inductive ind (mib,mip) = let realargs,_ = List.chop mip.mind_nrealdecls mip.mind_arity_ctxt in applist (mkIndU ind, - Context.Rel.to_extended_list mip.mind_nrealdecls mib.mind_params_ctxt - @ Context.Rel.to_extended_list 0 realargs) + Context.Rel.to_extended_list mkRel mip.mind_nrealdecls mib.mind_params_ctxt + @ Context.Rel.to_extended_list mkRel 0 realargs) let my_it_mkLambda_or_LetIn s c = it_mkLambda_or_LetIn c s let my_it_mkProd_or_LetIn s c = Term.it_mkProd_or_LetIn c s @@ -172,7 +172,7 @@ let build_sym_scheme env ind = let (mib,mip as specif),nrealargs,realsign,paramsctxt,paramsctxt1 = get_sym_eq_data env indu in let cstr n = - mkApp (mkConstructUi(indu,1),Context.Rel.to_extended_vect n mib.mind_params_ctxt) in + mkApp (mkConstructUi(indu,1),Context.Rel.to_extended_vect mkRel n mib.mind_params_ctxt) in let varH = fresh env (default_id_of_sort (snd (mind_arity mip))) in let applied_ind = build_dependent_inductive indu specif in let realsign_ind = @@ -185,7 +185,7 @@ let build_sym_scheme env ind = my_it_mkLambda_or_LetIn_name (lift_rel_context (nrealargs+1) realsign_ind) (mkApp (mkIndU indu,Array.concat - [Context.Rel.to_extended_vect (3*nrealargs+2) paramsctxt1; + [Context.Rel.to_extended_vect mkRel (3*nrealargs+2) paramsctxt1; rel_vect 1 nrealargs; rel_vect (2*nrealargs+2) nrealargs])), mkRel 1 (* varH *), @@ -226,13 +226,13 @@ let build_sym_involutive_scheme env ind = get_sym_eq_data env indu in let eq,eqrefl,ctx = get_coq_eq ctx in let sym, ctx, eff = const_of_scheme sym_scheme_kind env ind ctx in - let cstr n = mkApp (mkConstructUi (indu,1),Context.Rel.to_extended_vect n paramsctxt) in + let cstr n = mkApp (mkConstructUi (indu,1),Context.Rel.to_extended_vect mkRel n paramsctxt) in let varH = fresh env (default_id_of_sort (snd (mind_arity mip))) in let applied_ind = build_dependent_inductive indu specif in let applied_ind_C = mkApp (mkIndU indu, Array.append - (Context.Rel.to_extended_vect (nrealargs+1) mib.mind_params_ctxt) + (Context.Rel.to_extended_vect mkRel (nrealargs+1) mib.mind_params_ctxt) (rel_vect (nrealargs+1) nrealargs)) in let realsign_ind = name_context env ((LocalAssum (Name varH,applied_ind))::realsign) in @@ -246,15 +246,15 @@ let build_sym_involutive_scheme env ind = (mkApp (eq,[| mkApp (mkIndU indu, Array.concat - [Context.Rel.to_extended_vect (3*nrealargs+2) paramsctxt1; + [Context.Rel.to_extended_vect mkRel (3*nrealargs+2) paramsctxt1; rel_vect (2*nrealargs+2) nrealargs; rel_vect 1 nrealargs]); mkApp (sym,Array.concat - [Context.Rel.to_extended_vect (3*nrealargs+2) paramsctxt1; + [Context.Rel.to_extended_vect mkRel (3*nrealargs+2) paramsctxt1; rel_vect 1 nrealargs; rel_vect (2*nrealargs+2) nrealargs; [|mkApp (sym,Array.concat - [Context.Rel.to_extended_vect (3*nrealargs+2) paramsctxt1; + [Context.Rel.to_extended_vect mkRel (3*nrealargs+2) paramsctxt1; rel_vect (2*nrealargs+2) nrealargs; rel_vect 1 nrealargs; [|mkRel 1|]])|]]); @@ -337,7 +337,7 @@ let build_l2r_rew_scheme dep env ind kind = let eq,eqrefl,ctx = get_coq_eq ctx in let cstr n p = mkApp (mkConstructUi(indu,1), - Array.concat [Context.Rel.to_extended_vect n paramsctxt1; + Array.concat [Context.Rel.to_extended_vect mkRel n paramsctxt1; rel_vect p nrealargs]) in let varH = fresh env (default_id_of_sort (snd (mind_arity mip))) in let varHC = fresh env (Id.of_string "HC") in @@ -345,12 +345,12 @@ let build_l2r_rew_scheme dep env ind kind = let applied_ind = build_dependent_inductive indu specif in let applied_ind_P = mkApp (mkIndU indu, Array.concat - [Context.Rel.to_extended_vect (3*nrealargs) paramsctxt1; + [Context.Rel.to_extended_vect mkRel (3*nrealargs) paramsctxt1; rel_vect 0 nrealargs; rel_vect nrealargs nrealargs]) in let applied_ind_G = mkApp (mkIndU indu, Array.concat - [Context.Rel.to_extended_vect (3*nrealargs+3) paramsctxt1; + [Context.Rel.to_extended_vect mkRel (3*nrealargs+3) paramsctxt1; rel_vect (nrealargs+3) nrealargs; rel_vect 0 nrealargs]) in let realsign_P = lift_rel_context nrealargs realsign in @@ -361,10 +361,10 @@ let build_l2r_rew_scheme dep env ind kind = lift_rel_context (nrealargs+3) realsign) in let applied_sym_C n = mkApp(sym, - Array.append (Context.Rel.to_extended_vect n mip.mind_arity_ctxt) [|mkVar varH|]) in + Array.append (Context.Rel.to_extended_vect mkRel n mip.mind_arity_ctxt) [|mkVar varH|]) in let applied_sym_G = mkApp(sym, - Array.concat [Context.Rel.to_extended_vect (nrealargs*3+4) paramsctxt1; + Array.concat [Context.Rel.to_extended_vect mkRel (nrealargs*3+4) paramsctxt1; rel_vect (nrealargs+4) nrealargs; rel_vect 1 nrealargs; [|mkRel 1|]]) in @@ -374,7 +374,7 @@ let build_l2r_rew_scheme dep env ind kind = let ci = make_case_info (Global.env()) ind RegularStyle in let cieq = make_case_info (Global.env()) (fst (destInd eq)) RegularStyle in let applied_PC = - mkApp (mkVar varP,Array.append (Context.Rel.to_extended_vect 1 realsign) + mkApp (mkVar varP,Array.append (Context.Rel.to_extended_vect mkRel 1 realsign) (if dep then [|cstr (2*nrealargs+1) 1|] else [||])) in let applied_PG = mkApp (mkVar varP,Array.append (rel_vect 1 nrealargs) @@ -384,11 +384,11 @@ let build_l2r_rew_scheme dep env ind kind = (if dep then [|mkRel 2|] else [||])) in let applied_sym_sym = mkApp (sym,Array.concat - [Context.Rel.to_extended_vect (2*nrealargs+4) paramsctxt1; + [Context.Rel.to_extended_vect mkRel (2*nrealargs+4) paramsctxt1; rel_vect 4 nrealargs; rel_vect (nrealargs+4) nrealargs; [|mkApp (sym,Array.concat - [Context.Rel.to_extended_vect (2*nrealargs+4) paramsctxt1; + [Context.Rel.to_extended_vect mkRel (2*nrealargs+4) paramsctxt1; rel_vect (nrealargs+4) nrealargs; rel_vect 4 nrealargs; [|mkRel 2|]])|]]) in @@ -411,7 +411,7 @@ let build_l2r_rew_scheme dep env ind kind = mkApp (eq,[|lift 4 applied_ind;applied_sym_sym;mkRel 1|]), applied_PR)), mkApp (sym_involutive, - Array.append (Context.Rel.to_extended_vect 3 mip.mind_arity_ctxt) [|mkVar varH|]), + Array.append (Context.Rel.to_extended_vect mkRel 3 mip.mind_arity_ctxt) [|mkVar varH|]), [|main_body|]) else main_body)))))) @@ -450,7 +450,7 @@ let build_l2r_forward_rew_scheme dep env ind kind = get_sym_eq_data env indu in let cstr n p = mkApp (mkConstructUi(indu,1), - Array.concat [Context.Rel.to_extended_vect n paramsctxt1; + Array.concat [Context.Rel.to_extended_vect mkRel n paramsctxt1; rel_vect p nrealargs]) in let varH = fresh env (default_id_of_sort (snd (mind_arity mip))) in let varHC = fresh env (Id.of_string "HC") in @@ -458,12 +458,12 @@ let build_l2r_forward_rew_scheme dep env ind kind = let applied_ind = build_dependent_inductive indu specif in let applied_ind_P = mkApp (mkIndU indu, Array.concat - [Context.Rel.to_extended_vect (4*nrealargs+2) paramsctxt1; + [Context.Rel.to_extended_vect mkRel (4*nrealargs+2) paramsctxt1; rel_vect 0 nrealargs; rel_vect (nrealargs+1) nrealargs]) in let applied_ind_P' = mkApp (mkIndU indu, Array.concat - [Context.Rel.to_extended_vect (3*nrealargs+1) paramsctxt1; + [Context.Rel.to_extended_vect mkRel (3*nrealargs+1) paramsctxt1; rel_vect 0 nrealargs; rel_vect (2*nrealargs+1) nrealargs]) in let realsign_P n = lift_rel_context (nrealargs*n+n) realsign in @@ -541,7 +541,7 @@ let build_r2l_forward_rew_scheme dep env ind kind = let ((mib,mip as specif),constrargs,realsign,paramsctxt,nrealargs) = get_non_sym_eq_data env indu in let cstr n = - mkApp (mkConstructUi(indu,1),Context.Rel.to_extended_vect n mib.mind_params_ctxt) in + mkApp (mkConstructUi(indu,1),Context.Rel.to_extended_vect mkRel n mib.mind_params_ctxt) in let constrargs_cstr = constrargs@[cstr 0] in let varH = fresh env (default_id_of_sort (snd (mind_arity mip))) in let varHC = fresh env (Id.of_string "HC") in @@ -557,8 +557,8 @@ let build_r2l_forward_rew_scheme dep env ind kind = applist (mkVar varP,if dep then constrargs_cstr else constrargs) in let applied_PG = mkApp (mkVar varP, - if dep then Context.Rel.to_extended_vect 0 realsign_ind - else Context.Rel.to_extended_vect 1 realsign) in + if dep then Context.Rel.to_extended_vect mkRel 0 realsign_ind + else Context.Rel.to_extended_vect mkRel 1 realsign) in let c = (my_it_mkLambda_or_LetIn paramsctxt (my_it_mkLambda_or_LetIn_name realsign_ind @@ -608,7 +608,7 @@ let fix_r2l_forward_rew_scheme (c, ctx') = (mkLambda_or_LetIn (RelDecl.map_constr (lift 2) ind) (EConstr.Unsafe.to_constr (Reductionops.whd_beta Evd.empty (EConstr.of_constr (applist (c, - Context.Rel.to_extended_list 3 indargs @ [mkRel 1;mkRel 3;mkRel 2])))))))) + Context.Rel.to_extended_list mkRel 3 indargs @ [mkRel 1;mkRel 3;mkRel 2])))))))) in c', ctx' | _ -> anomaly (Pp.str "Ill-formed non-dependent left-to-right rewriting scheme") @@ -763,8 +763,8 @@ let build_congr env (eq,refl,ctx) ind = (mkNamedLambda varH (applist (mkIndU indu, - Context.Rel.to_extended_list (mip.mind_nrealargs+2) paramsctxt @ - Context.Rel.to_extended_list 0 realsign)) + Context.Rel.to_extended_list mkRel (mip.mind_nrealargs+2) paramsctxt @ + Context.Rel.to_extended_list mkRel 0 realsign)) (mkCase (ci, my_it_mkLambda_or_LetIn_name (lift_rel_context (mip.mind_nrealargs+3) realsign) @@ -772,9 +772,9 @@ let build_congr env (eq,refl,ctx) ind = (Anonymous, applist (mkIndU indu, - Context.Rel.to_extended_list (2*mip.mind_nrealdecls+3) + Context.Rel.to_extended_list mkRel (2*mip.mind_nrealdecls+3) paramsctxt - @ Context.Rel.to_extended_list 0 realsign), + @ Context.Rel.to_extended_list mkRel 0 realsign), mkApp (eq, [|mkVar varB; mkApp (mkVar varf, [|lift (2*mip.mind_nrealdecls+4) b|]); diff --git a/tactics/tactics.ml b/tactics/tactics.ml index 0ecccd5c0..8260c14ad 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -1634,7 +1634,7 @@ let make_projection env sigma params cstr sign elim i n c u = then let t = lift (i+1-n) t in let abselim = beta_applist sigma (elim, params@[t;branch]) in - let args = Array.map EConstr.of_constr (Context.Rel.to_extended_vect 0 sign) in + let args = Context.Rel.to_extended_vect mkRel 0 sign in let c = beta_applist sigma (abselim, [mkApp (c, args)]) in Some (it_mkLambda_or_LetIn c sign, it_mkProd_or_LetIn t sign) else @@ -1643,8 +1643,7 @@ let make_projection env sigma params cstr sign elim i n c u = (* goes from left to right when i increases! *) match List.nth l i with | Some proj -> - let args = Context.Rel.to_extended_vect 0 sign in - let args = Array.map EConstr.of_constr args in + let args = Context.Rel.to_extended_vect mkRel 0 sign in let proj = if Environ.is_projection proj env then mkProj (Projection.make proj false, mkApp (c, args)) @@ -2190,7 +2189,7 @@ let bring_hyps hyps = let store = Proofview.Goal.extra gl in let concl = Tacmach.New.pf_nf_concl gl in let newcl = List.fold_right mkNamedProd_or_LetIn hyps concl in - let args = Array.map_of_list EConstr.of_constr (Context.Named.to_instance hyps) in + let args = Array.of_list (Context.Named.to_instance mkVar hyps) in Refine.refine { run = begin fun sigma -> let Sigma (ev, sigma, p) = Evarutil.new_evar env sigma ~principal:true ~store newcl in @@ -2868,8 +2867,7 @@ let old_generalize_dep ?(with_let=false) c gl = (cl',project gl) in (** Check that the generalization is indeed well-typed *) let (evd, _) = Typing.type_of env evd cl'' in - let args = Context.Named.to_instance to_quantify_rev in - let args = List.map EConstr.of_constr args in + let args = Context.Named.to_instance mkVar to_quantify_rev in tclTHENLIST [tclEVARS evd; Proofview.V82.of_tactic (apply_type cl'' (if Option.is_empty body then c::args else args)); @@ -3994,7 +3992,7 @@ let compute_scheme_signature evd scheme names_info ind_type_guess = let ind_is_ok = List.equal (fun c1 c2 -> EConstr.eq_constr evd c1 c2) (List.lastn scheme.nargs indargs) - (List.map EConstr.of_constr (Context.Rel.to_extended_list 0 scheme.args)) in + (Context.Rel.to_extended_list mkRel 0 scheme.args) in if not (ccl_arg_ok && ind_is_ok) then error_ind_scheme "the conclusion of" in (cond, check_concl) @@ -4965,7 +4963,7 @@ let abstract_subproof id gk tac = in let const, args = if !shrink_abstract then shrink_entry sign const - else (const, List.rev (Context.Named.to_instance sign)) + else (const, List.rev (Context.Named.to_instance Constr.mkVar sign)) in let args = List.map EConstr.of_constr args in let cd = Entries.DefinitionEntry const in diff --git a/toplevel/auto_ind_decl.ml b/toplevel/auto_ind_decl.ml index 14071d869..f343cc73d 100644 --- a/toplevel/auto_ind_decl.ml +++ b/toplevel/auto_ind_decl.ml @@ -104,7 +104,7 @@ let mkFullInd (ind,u) n = context_chop (nparams-nparrec) mib.mind_params_ctxt in if nparrec > 0 then mkApp (mkIndU (ind,u), - Array.of_list(Context.Rel.to_extended_list (nparrec+n) lnamesparrec)) + Array.of_list(Context.Rel.to_extended_list mkRel (nparrec+n) lnamesparrec)) else mkIndU (ind,u) let check_bool_is_defined () = @@ -139,7 +139,7 @@ let build_beq_scheme mode kn = | Name s -> Id.of_string ("eq_"^(Id.to_string s)) | Anonymous -> Id.of_string "eq_A" in - let ext_rel_list = Context.Rel.to_extended_list 0 lnamesparrec in + let ext_rel_list = Context.Rel.to_extended_list mkRel 0 lnamesparrec in let lift_cnt = ref 0 in let eqs_typ = List.map (fun aa -> let a = lift !lift_cnt aa in @@ -241,7 +241,7 @@ let build_beq_scheme mode kn = Cn => match Y with ... end |] part *) let ci = make_case_info env (fst ind) MatchStyle in let constrs n = get_constructors env (make_ind_family (ind, - Context.Rel.to_extended_list (n+nb_ind-1) mib.mind_params_ctxt)) in + Context.Rel.to_extended_list mkRel (n+nb_ind-1) mib.mind_params_ctxt)) in let constrsi = constrs (3+nparrec) in let n = Array.length constrsi in let ar = Array.make n (Lazy.force ff) in diff --git a/toplevel/class.ml b/toplevel/class.ml index 4c7aa01cd..95114ec39 100644 --- a/toplevel/class.ml +++ b/toplevel/class.ml @@ -194,13 +194,13 @@ let build_id_coercion idf_opt source poly = let val_f = it_mkLambda_or_LetIn (mkLambda (Name Namegen.default_dependent_ident, - applistc vs (Context.Rel.to_extended_list 0 lams), + applistc vs (Context.Rel.to_extended_list mkRel 0 lams), mkRel 1)) lams in let typ_f = List.fold_left (fun d c -> Term.mkProd_wo_LetIn c d) - (mkProd (Anonymous, applistc vs (Context.Rel.to_extended_list 0 lams), lift 1 t)) + (mkProd (Anonymous, applistc vs (Context.Rel.to_extended_list mkRel 0 lams), lift 1 t)) lams in (* juste pour verification *) diff --git a/toplevel/discharge.ml b/toplevel/discharge.ml index e24d5e74f..b898f3e83 100644 --- a/toplevel/discharge.ml +++ b/toplevel/discharge.ml @@ -39,7 +39,7 @@ let detype_param = let abstract_inductive hyps nparams inds = let ntyp = List.length inds in let nhyp = Context.Named.length hyps in - let args = Context.Named.to_instance (List.rev hyps) in + let args = Context.Named.to_instance mkVar (List.rev hyps) in let args = Array.of_list args in let subs = List.init ntyp (fun k -> lift nhyp (mkApp(mkRel (k+1),args))) in let inds' = diff --git a/toplevel/record.ml b/toplevel/record.ml index a65f5252e..ec719bca8 100644 --- a/toplevel/record.ml +++ b/toplevel/record.ml @@ -273,8 +273,8 @@ let declare_projections indsp ?(kind=StructureComponent) binder_name coers field let ctx = Univ.instantiate_univ_context mib.mind_universes in let indu = indsp, u in let r = mkIndU (indsp,u) in - let rp = applist (r, Context.Rel.to_extended_list 0 paramdecls) in - let paramargs = Context.Rel.to_extended_list 1 paramdecls in (*def in [[params;x:rp]]*) + let rp = applist (r, Context.Rel.to_extended_list mkRel 0 paramdecls) in + let paramargs = Context.Rel.to_extended_list mkRel 1 paramdecls in (*def in [[params;x:rp]]*) let x = Name binder_name in let fields = instantiate_possibly_recursive_type indu paramdecls fields in let lifted_fields = Termops.lift_rel_context 1 fields in @@ -386,7 +386,7 @@ open Typeclasses let declare_structure finite poly ctx id idbuild paramimpls params arity template fieldimpls fields ?(kind=StructureComponent) ?name is_coe coers sign = let nparams = List.length params and nfields = List.length fields in - let args = Context.Rel.to_extended_list nfields params in + let args = Context.Rel.to_extended_list mkRel nfields params in let ind = applist (mkRel (1+nparams+nfields), args) in let type_constructor = it_mkProd_or_LetIn ind fields in let binder_name = -- cgit v1.2.3 From b4b90c5d2e8c413e1981c456c933f35679386f09 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sat, 26 Nov 2016 16:18:47 +0100 Subject: Definining EConstr-based contexts. This removes quite a few unsafe casts. Unluckily, I had to reintroduce the old non-module based names for these data structures, because I could not reproduce easily the same hierarchy in EConstr. --- engine/eConstr.ml | 105 +++++++++++----- engine/eConstr.mli | 48 ++++++-- engine/evarutil.ml | 18 +-- engine/evarutil.mli | 99 +++++++-------- engine/proofview.ml | 2 +- engine/proofview.mli | 15 +-- engine/termops.ml | 94 ++++++++------- engine/termops.mli | 49 ++++---- interp/constrintern.ml | 4 +- interp/constrintern.mli | 4 +- ltac/rewrite.ml | 43 ++----- ltac/rewrite.mli | 4 +- ltac/tactic_matching.ml | 4 +- ltac/tactic_matching.mli | 2 +- plugins/cc/cctac.ml | 8 +- plugins/decl_mode/decl_proof_instr.ml | 15 +-- plugins/extraction/extraction.ml | 2 +- plugins/firstorder/instances.ml | 3 +- plugins/firstorder/rules.ml | 2 +- plugins/fourier/fourierR.ml | 4 +- plugins/funind/functional_principles_proofs.ml | 7 +- plugins/funind/functional_principles_types.ml | 16 +-- plugins/funind/indfun.ml | 2 +- plugins/funind/invfun.ml | 20 ++- plugins/funind/merge.ml | 2 +- plugins/funind/recdef.ml | 2 +- plugins/omega/coq_omega.ml | 30 ++--- plugins/rtauto/refl_tauto.ml | 28 +++-- plugins/rtauto/refl_tauto.mli | 6 +- plugins/ssrmatching/ssrmatching.ml4 | 4 +- pretyping/cases.ml | 100 ++++++++------- pretyping/cases.mli | 8 +- pretyping/coercion.ml | 14 +-- pretyping/constr_matching.ml | 26 ++-- pretyping/evarconv.ml | 16 +-- pretyping/evardefine.ml | 15 +-- pretyping/evarsolve.ml | 27 ++--- pretyping/find_subterm.ml | 4 +- pretyping/find_subterm.mli | 4 +- pretyping/inductiveops.ml | 2 +- pretyping/patternops.ml | 17 +-- pretyping/pretyping.ml | 52 ++++---- pretyping/reductionops.ml | 40 +++--- pretyping/reductionops.mli | 6 +- pretyping/retyping.ml | 24 ++-- pretyping/retyping.mli | 2 +- pretyping/tacred.ml | 55 ++++----- pretyping/typeclasses.ml | 1 + pretyping/typeclasses.mli | 2 +- pretyping/typing.ml | 21 ++-- pretyping/unification.ml | 8 +- pretyping/unification.mli | 2 +- printing/prettyp.ml | 8 +- proofs/goal.ml | 1 + proofs/logic.ml | 21 ++-- proofs/logic.mli | 2 +- proofs/refiner.ml | 6 +- proofs/refiner.mli | 3 +- proofs/tacmach.ml | 6 +- proofs/tacmach.mli | 10 +- tactics/auto.ml | 2 +- tactics/class_tactics.ml | 16 +-- tactics/contradiction.ml | 4 +- tactics/equality.ml | 30 ++--- tactics/hints.ml | 4 +- tactics/hints.mli | 4 +- tactics/hipattern.ml | 5 +- tactics/inv.ml | 7 +- tactics/leminv.ml | 20 ++- tactics/tacticals.ml | 3 +- tactics/tacticals.mli | 28 ++--- tactics/tactics.ml | 161 ++++++++++--------------- tactics/tactics.mli | 18 +-- tactics/term_dnet.ml | 5 +- toplevel/classes.ml | 2 + toplevel/command.ml | 56 +++++---- toplevel/himsg.ml | 62 ++++------ toplevel/record.ml | 1 + toplevel/vernacentries.ml | 2 +- 79 files changed, 759 insertions(+), 816 deletions(-) diff --git a/engine/eConstr.ml b/engine/eConstr.ml index 629be8e4b..657285de2 100644 --- a/engine/eConstr.ml +++ b/engine/eConstr.ml @@ -77,6 +77,10 @@ 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, Univ.Instance.empty) @@ -201,14 +205,6 @@ let decompose_app sigma c = | App (f,cl) -> (f, Array.to_list cl) | _ -> (c,[]) -let local_assum (na, t) = - let open Rel.Declaration in - LocalAssum (na, unsafe_to_constr t) - -let local_def (na, b, t) = - let open Rel.Declaration in - LocalDef (na, unsafe_to_constr b, unsafe_to_constr t) - 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 @@ -218,38 +214,41 @@ let decompose_lam sigma c = 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 (local_assum (x,t)) l) c - | LetIn (x,b,t,c) -> lamdec_rec (Context.Rel.add (local_def (x,b,t)) l) c + | 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 error "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 (local_assum (x,t)) l) (n-1) c - | LetIn (x,b,t,c) -> lamdec_rec (Context.Rel.add (local_def (x,b,t)) l) n c + | 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 -> error "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 error "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 (local_assum (x,t)) l) (n-1) c - | LetIn (x,b,t,c) -> lamdec_rec (Context.Rel.add (local_def (x,b,t)) l) (n-1) c + | 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 -> error "decompose_lam_n_decls: not enough abstractions" in @@ -264,24 +263,26 @@ let decompose_prod sigma c = 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 (local_assum (x,t)) l) c - | LetIn (x,b,t,c) -> proddec_rec (Context.Rel.add (local_def (x,b,t)) l) c + | 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 error "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 (local_assum (x,t)) l) (n-1) c - | LetIn (x,b,t,c) -> prodec_rec (Context.Rel.add (local_def (x,b,t)) l) (n-1) c + | 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 -> error "decompose_prod_n_assum: not enough assumptions" in @@ -413,24 +414,26 @@ let iter sigma f c = match kind sigma c with | 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 = match kind sigma c with +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 (local_assum (na, t)) n) c - | Lambda (na,t,c) -> f n t; f (g (local_assum (na, t)) n) c - | LetIn (na,b,t,c) -> f n b; f n t; f (g (local_def (na, b, t)) n) c + | 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 (local_assum (na,t)) n) n lna tl in + 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 (local_assum (na,t)) n) n lna tl in + 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 = @@ -611,14 +614,14 @@ let rec isArity sigma c = let mkProd_or_LetIn decl c = let open Context.Rel.Declaration in match decl with - | LocalAssum (na,t) -> mkProd (na, of_constr t, c) - | LocalDef (na,b,t) -> mkLetIn (na, of_constr b, of_constr t, c) + | 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, of_constr t, c) - | LocalDef (na,b,t) -> mkLetIn (na, of_constr b, of_constr t, c) + | 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) @@ -627,18 +630,56 @@ 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 (of_constr t) c - | LocalDef (id,b,t) -> mkNamedLetIn id (of_constr b) (of_constr t) c + | 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 (of_constr t) c - | LocalDef (id,b,t) -> mkNamedLetIn id (of_constr b) (of_constr t) c + | 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 +open Context +open Environ + +let sym : type a b. (a, b) eq -> (b, a) eq = fun Refl -> Refl + +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 + +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) + + module Unsafe = struct let to_constr = unsafe_to_constr diff --git a/engine/eConstr.mli b/engine/eConstr.mli index cb671154c..d6b2113d2 100644 --- a/engine/eConstr.mli +++ b/engine/eConstr.mli @@ -10,6 +10,7 @@ open CSig open Names open Constr open Context +open Environ type t (** Type of incomplete terms. Essentially a wrapper around {!Constr.t} ensuring @@ -22,6 +23,10 @@ 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 Destructors} *) @@ -77,16 +82,16 @@ val mkArrow : t -> t -> t val applist : t * t list -> t -val mkProd_or_LetIn : Rel.Declaration.t -> t -> t -val mkLambda_or_LetIn : Rel.Declaration.t -> t -> t -val it_mkProd_or_LetIn : t -> Rel.t -> t -val it_mkLambda_or_LetIn : t -> Rel.t -> 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.t -> types -> types -val mkNamedProd_or_LetIn : Named.Declaration.t -> types -> types +val mkNamedLambda_or_LetIn : named_declaration -> types -> types +val mkNamedProd_or_LetIn : named_declaration -> types -> types (** {6 Simple case analysis} *) @@ -131,13 +136,13 @@ 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 -> Context.Rel.t * t -val decompose_lam_n_assum : Evd.evar_map -> int -> t -> Context.Rel.t * t -val decompose_lam_n_decls : Evd.evar_map -> int -> t -> Context.Rel.t * 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 decompose_prod : Evd.evar_map -> t -> (Name.t * t) list * t -val decompose_prod_assum : Evd.evar_map -> t -> Context.Rel.t * t -val decompose_prod_n_assum : Evd.evar_map -> int -> t -> Context.Rel.t * 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 @@ -156,7 +161,7 @@ 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.t -> '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} *) @@ -184,6 +189,25 @@ val subst_univs_level_constr : Univ.universe_level_subst -> t -> t 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 + (** {5 Unsafe operations} *) module Unsafe : diff --git a/engine/evarutil.ml b/engine/evarutil.ml index 42620c0b5..0582e34be 100644 --- a/engine/evarutil.ml +++ b/engine/evarutil.ml @@ -139,12 +139,12 @@ let nf_named_context_evar sigma ctx = Context.Named.map (nf_evar0 sigma) ctx let nf_rel_context_evar sigma ctx = - Context.Rel.map (nf_evar0 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_evar0 evc) info @@ -320,7 +320,7 @@ let empty_csubst = (0, Int.Map.empty) type ext_named_context = csubst * (Id.t * EConstr.constr) list * - Id.Set.t * Context.Named.t + Id.Set.t * EConstr.named_context let push_var id (n, s) = let s = Int.Map.add n (EConstr.mkVar id) s in @@ -330,7 +330,7 @@ let push_rel_decl_to_named_context decl (subst, vsubst, avoid, nc) = let open EConstr in let open Vars in let map_decl f d = - NamedDecl.map_constr (fun c -> EConstr.Unsafe.to_constr (f (EConstr.of_constr c))) d + NamedDecl.map_constr f d in let replace_var_named_declaration id0 id decl = let id' = NamedDecl.get_id decl in @@ -354,7 +354,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 (RelDecl.get_type decl) na) avoid + next_ident_away (id_of_name_using_hdchar empty_env (EConstr.Unsafe.to_constr (RelDecl.get_type decl)) na) avoid in match extract_if_neq id na with | Some id0 when not (is_section_variable id0) -> @@ -542,6 +542,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) -> @@ -595,6 +596,7 @@ 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 @@ -616,10 +618,9 @@ 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 = - let concl = EConstr.Unsafe.to_constr concl in match clear_hyps_in_evi_main env evdref hyps [concl] ids with | (nhyps,[nconcl]) -> (nhyps,nconcl) | _ -> assert false @@ -746,6 +747,7 @@ 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 open EConstr in 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) diff --git a/engine/evarutil.mli b/engine/evarutil.mli index 67de18abc..2da4f8043 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. *) @@ -17,51 +18,51 @@ open Environ (** [new_meta] is a generator of unique meta variables *) val new_meta : unit -> metavariable -val mk_new_meta : unit -> EConstr.constr +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 -> - ?candidates:EConstr.constr list -> ?store:Store.t -> + ?candidates:constr list -> ?store:Store.t -> ?naming:Misctypes.intro_pattern_naming_expr -> - ?principal:bool -> EConstr.types -> (EConstr.constr, 'r) Sigma.sigma + ?principal:bool -> types -> (constr, 'r) Sigma.sigma val new_pure_evar : named_context_val -> 'r Sigma.t -> ?src:Loc.t * Evar_kinds.t -> ?filter:Filter.t -> - ?candidates:EConstr.constr list -> ?store:Store.t -> + ?candidates:constr list -> ?store:Store.t -> ?naming:Misctypes.intro_pattern_naming_expr -> - ?principal:bool -> EConstr.types -> (evar, 'r) Sigma.sigma + ?principal:bool -> types -> (evar, 'r) Sigma.sigma val new_pure_evar_full : 'r Sigma.t -> evar_info -> (evar, 'r) Sigma.sigma (** the same with side-effects *) val e_new_evar : env -> evar_map ref -> ?src:Loc.t * Evar_kinds.t -> ?filter:Filter.t -> - ?candidates:EConstr.constr list -> ?store:Store.t -> + ?candidates:constr list -> ?store:Store.t -> ?naming:Misctypes.intro_pattern_naming_expr -> - ?principal:bool -> EConstr.types -> EConstr.constr + ?principal:bool -> types -> constr (** 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 -> ?naming:Misctypes.intro_pattern_naming_expr -> ?principal:bool -> rigid -> - (EConstr.constr * sorts, 'r) Sigma.sigma + (constr * sorts, 'r) Sigma.sigma val e_new_type_evar : env -> evar_map ref -> ?src:Loc.t * Evar_kinds.t -> ?filter:Filter.t -> - ?naming:Misctypes.intro_pattern_naming_expr -> ?principal:bool -> rigid -> EConstr.constr * sorts + ?naming:Misctypes.intro_pattern_naming_expr -> ?principal:bool -> rigid -> constr * sorts -val new_Type : ?rigid:rigid -> env -> 'r Sigma.t -> (EConstr.constr, 'r) Sigma.sigma -val e_new_Type : ?rigid:rigid -> env -> evar_map ref -> EConstr.constr +val new_Type : ?rigid:rigid -> env -> 'r Sigma.t -> (constr, 'r) Sigma.sigma +val e_new_Type : ?rigid:rigid -> env -> evar_map ref -> constr val restrict_evar : 'r Sigma.t -> existential_key -> Filter.t -> - EConstr.constr list option -> (existential_key, 'r) Sigma.sigma + constr list option -> (existential_key, 'r) Sigma.sigma (** Polymorphic constants *) -val new_global : 'r Sigma.t -> Globnames.global_reference -> (EConstr.constr, 'r) Sigma.sigma -val e_new_global : evar_map ref -> Globnames.global_reference -> EConstr.constr +val new_global : 'r Sigma.t -> Globnames.global_reference -> (constr, 'r) Sigma.sigma +val e_new_global : evar_map ref -> Globnames.global_reference -> constr (** Create a fresh evar in a context different from its definition context: [new_evar_instance sign evd ty inst] creates a new evar of context @@ -70,15 +71,15 @@ val e_new_global : evar_map ref -> Globnames.global_reference -> EConstr.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 -> EConstr.types -> - ?src:Loc.t * Evar_kinds.t -> ?filter:Filter.t -> ?candidates:EConstr.constr list -> + named_context_val -> 'r Sigma.t -> types -> + ?src:Loc.t * Evar_kinds.t -> ?filter:Filter.t -> ?candidates:constr list -> ?store:Store.t -> ?naming:Misctypes.intro_pattern_naming_expr -> ?principal:bool -> - EConstr.constr list -> (EConstr.constr, 'r) Sigma.sigma + constr list -> (constr, 'r) Sigma.sigma 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,15 +89,15 @@ 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 : evar_map -> EConstr.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 -> EConstr.constr -> EConstr.constr +val whd_head_evar : evar_map -> constr -> constr (* An over-approximation of [has_undefined (nf_evars evd c)] *) -val has_undefined_evars : evar_map -> EConstr.constr -> bool +val has_undefined_evars : evar_map -> constr -> bool -val is_ground_term : evar_map -> EConstr.constr -> bool +val is_ground_term : evar_map -> constr -> bool val is_ground_env : evar_map -> env -> bool (** [gather_dependent_evars evm seeds] classifies the evars in [evm] @@ -121,14 +122,14 @@ val advance : evar_map -> evar -> evar option This is roughly a combination of the previous functions and [nf_evar]. *) -val undefined_evars_of_term : evar_map -> EConstr.constr -> Evar.Set.t +val undefined_evars_of_term : evar_map -> constr -> Evar.Set.t val undefined_evars_of_named_context : evar_map -> Context.Named.t -> Evar.Set.t 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 -> EConstr.t -> bool +val occur_evar_upto : evar_map -> Evar.t -> constr -> bool (** {6 Value/Type constraints} *) @@ -139,18 +140,18 @@ val judge_of_new_Type : 'r Sigma.t -> (unsafe_judgment, 'r) Sigma.sigma (** [flush_and_check_evars] raise [Uninstantiated_evar] if an evar remains uninstantiated; [nf_evar] leaves uninstantiated evars as is *) -val whd_evar : evar_map -> EConstr.constr -> EConstr.constr -val nf_evar : evar_map -> EConstr.constr -> EConstr.constr -val j_nf_evar : evar_map -> EConstr.unsafe_judgment -> EConstr.unsafe_judgment +val whd_evar : evar_map -> constr -> constr +val nf_evar : evar_map -> constr -> constr +val j_nf_evar : evar_map -> unsafe_judgment -> unsafe_judgment val jl_nf_evar : - evar_map -> EConstr.unsafe_judgment list -> EConstr.unsafe_judgment list + evar_map -> unsafe_judgment list -> unsafe_judgment list val jv_nf_evar : - evar_map -> EConstr.unsafe_judgment array -> EConstr.unsafe_judgment array + evar_map -> unsafe_judgment array -> unsafe_judgment array val tj_nf_evar : - evar_map -> EConstr.unsafe_type_judgment -> EConstr.unsafe_type_judgment + 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,39 @@ 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 -> EConstr.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) 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 @@ -199,7 +200,7 @@ exception ClearDependencyError of Id.t * clear_dependency_error used by [Goal] and (indirectly) [Proofview] to handle the clear tactic gracefully*) val cleared : bool Store.field -val clear_hyps_in_evi : env -> evar_map ref -> named_context_val -> EConstr.types -> +val clear_hyps_in_evi : env -> evar_map ref -> named_context_val -> types -> Id.Set.t -> named_context_val * types val clear_hyps2_in_evi : env -> evar_map ref -> named_context_val -> types -> types -> @@ -208,19 +209,19 @@ val clear_hyps2_in_evi : env -> evar_map ref -> named_context_val -> types -> ty type csubst val empty_csubst : csubst -val csubst_subst : csubst -> EConstr.t -> EConstr.t +val csubst_subst : csubst -> constr -> constr type ext_named_context = - csubst * (Id.t * EConstr.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 + rel_declaration -> ext_named_context -> ext_named_context -val push_rel_context_to_named_context : Environ.env -> EConstr.types -> - named_context_val * EConstr.types * EConstr.constr list * csubst * (identifier*EConstr.constr) list +val push_rel_context_to_named_context : Environ.env -> types -> + named_context_val * types * constr list * csubst * (identifier*constr) list -val generalize_evar_over_rels : evar_map -> EConstr.existential -> EConstr.types * EConstr.constr list +val generalize_evar_over_rels : evar_map -> existential -> types * constr list (** Evar combinators *) @@ -235,5 +236,5 @@ val meta_counter_summary_name : string (** Deprecater *) -type type_constraint = EConstr.types option -type val_constraint = EConstr.constr option +type type_constraint = types option +type val_constraint = constr option diff --git a/engine/proofview.ml b/engine/proofview.ml index ab72cc405..0a18cf191 100644 --- a/engine/proofview.ml +++ b/engine/proofview.ml @@ -995,7 +995,7 @@ module Goal = struct let env { env=env } = env let sigma { sigma=sigma } = Sigma.Unsafe.of_evar_map sigma - let hyps { env=env } = Environ.named_context env + let hyps { env=env } = EConstr.named_context env let concl { concl=concl } = concl let extra { sigma=sigma; self=self } = goal_extra sigma self diff --git a/engine/proofview.mli b/engine/proofview.mli index 9f10e874a..025e3de20 100644 --- a/engine/proofview.mli +++ b/engine/proofview.mli @@ -14,6 +14,7 @@ open Util open Term +open EConstr (** Main state of tactics *) type proofview @@ -43,7 +44,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 * EConstr.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 +53,7 @@ val init : Evd.evar_map -> (Environ.env * EConstr.types) list -> entry * proofvi [evar_map] is threaded in state passing style. *) type telescope = | TNil of Evd.evar_map - | TCons of Environ.env * Evd.evar_map * EConstr.types * (Evd.evar_map -> EConstr.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 @@ -69,8 +70,8 @@ val finished : proofview -> bool (** Returns the current [evar] state. *) val return : proofview -> Evd.evar_map -val partial_proof : entry -> proofview -> EConstr.constr list -val initial_goals : entry -> (EConstr.constr * EConstr.types) list +val partial_proof : entry -> proofview -> constr list +val initial_goals : entry -> (constr * types) list @@ -484,15 +485,15 @@ module Goal : sig 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 -> EConstr.constr - val hyps : ([ `NF ], 'r) t -> Context.Named.t + val concl : ([ `NF ], 'r) t -> constr + val hyps : ([ `NF ], 'r) t -> named_context 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 -> EConstr.constr + val raw_concl : ('a, 'r) t -> constr type ('a, 'b) enter = { enter : 'r. ('a, 'r) t -> 'b } diff --git a/engine/termops.ml b/engine/termops.ml index 46e9ad927..c35e92f97 100644 --- a/engine/termops.ml +++ b/engine/termops.ml @@ -167,19 +167,10 @@ let rel_list n m = in reln [] 1 -let local_assum (na, t) = - let open RelDecl in - let inj = EConstr.Unsafe.to_constr in - LocalAssum (na, inj t) - -let local_def (na, b, t) = - let open RelDecl in - let inj = EConstr.Unsafe.to_constr in - LocalDef (na, inj b, inj t) - let push_rel_assum (x,t) env = let open RelDecl in - push_rel (local_assum (x,t)) env + let open EConstr in + push_rel (LocalAssum (x,t)) env let push_rels_assum assums = let open RelDecl in @@ -218,8 +209,8 @@ let mkProd_wo_LetIn decl c = let open EConstr in let open RelDecl in match decl with - | LocalAssum (na,t) -> mkProd (na, EConstr.of_constr t, c) - | LocalDef (_,b,_) -> Vars.subst1 (EConstr.of_constr b) c + | LocalAssum (na,t) -> mkProd (na, t, c) + | LocalDef (_,b,_) -> Vars.subst1 b c 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 @@ -334,7 +325,8 @@ let map_constr_with_named_binders g f l c = match kind_of_term c with let fold_rec_types g (lna,typarray,_) e = let open EConstr in - let ctxt = Array.map2_i (fun i na t -> local_assum (na, Vars.lift i t)) lna typarray 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 let map_left2 f a g b = @@ -350,6 +342,7 @@ let map_left2 f a g b = end let map_constr_with_binders_left_to_right sigma g f l c = + let open RelDecl in let open EConstr in match EConstr.kind sigma c with | (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _ @@ -361,18 +354,18 @@ let map_constr_with_binders_left_to_right sigma g f l c = else mkCast (b',k,t') | Prod (na,t,b) -> let t' = f l t in - let b' = f (g (local_assum (na,t)) l) b in + let b' = f (g (LocalAssum (na,t)) l) b in if t' == t && b' == b then c else mkProd (na, t', b') | Lambda (na,t,b) -> let t' = f l t in - let b' = f (g (local_assum (na,t)) l) b in + let b' = f (g (LocalAssum (na,t)) l) b in if t' == t && b' == b then c else mkLambda (na, t', b') | LetIn (na,bo,t,b) -> let bo' = f l bo in let t' = f l t in - let b' = f (g (local_def (na,bo,t)) l) b in + let b' = f (g (LocalDef (na,bo,t)) l) b in if bo' == bo && t' == t && b' == b then c else mkLetIn (na, bo', t', b') | App (c,[||]) -> assert false @@ -414,7 +407,6 @@ let map_constr_with_binders_left_to_right sigma g f l c = (* strong *) let map_constr_with_full_binders sigma g f l cstr = - let inj c = EConstr.Unsafe.to_constr c in let open EConstr in let open RelDecl in match EConstr.kind sigma cstr with @@ -426,16 +418,16 @@ let map_constr_with_full_binders sigma 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, inj 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, inj 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, inj b, inj 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 @@ -456,7 +448,7 @@ let map_constr_with_full_binders sigma 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, inj 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 @@ -464,7 +456,7 @@ let map_constr_with_full_binders sigma 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, inj 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 @@ -577,10 +569,10 @@ let occur_var env sigma id c = let occur_var_in_decl env sigma hyp decl = let open NamedDecl in match decl with - | LocalAssum (_,typ) -> occur_var env sigma hyp (EConstr.of_constr typ) + | LocalAssum (_,typ) -> occur_var env sigma hyp typ | LocalDef (_, body, typ) -> - occur_var env sigma hyp (EConstr.of_constr typ) || - occur_var env sigma hyp (EConstr.of_constr body) + occur_var env sigma hyp typ || + occur_var env sigma hyp body let local_occur_var sigma id c = let rec occur c = match EConstr.kind sigma c with @@ -655,8 +647,8 @@ let dependent_univs_no_evar sigma c t = dependent_main true true sigma c t let dependent_in_decl sigma a decl = let open NamedDecl in match decl with - | LocalAssum (_,t) -> dependent sigma a (EConstr.of_constr t) - | LocalDef (_, body, t) -> dependent sigma a (EConstr.of_constr body) || dependent sigma a (EConstr.of_constr t) + | LocalAssum (_,t) -> dependent sigma a t + | LocalDef (_, body, t) -> dependent sigma a body || dependent sigma a t let count_occurrences sigma m t = let open EConstr in @@ -886,7 +878,7 @@ 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 @@ -895,28 +887,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 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 sigma cv_pb (EConstr.of_constr c1) (EConstr.of_constr (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) -> @@ -925,21 +919,20 @@ let filtering sigma env cv_pb c1 c2 = | _, Evar (ev,_) -> define cv_pb env ev c1 | Evar (ev,_), _ -> define cv_pb env ev c2 | _ -> - let inj = EConstr.Unsafe.to_constr in if compare_constr_univ sigma - (fun pb c1 c2 -> aux env pb (inj c1) (inj c2); true) cv_pb (EConstr.of_constr c1) (EConstr.of_constr c2) then () + (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 sigma c = - let rec prodec_rec i l sigma c = match EConstr.kind sigma c with - | Prod (n,t,c) -> prodec_rec (succ i) (local_assum (n, t)::l) sigma c - | LetIn (n,d,t,c) -> prodec_rec (succ i) (local_def (n, d, t)::l) sigma c - | Cast (c,_,_) -> prodec_rec i l sigma c - | _ -> i,l, c in - prodec_rec 0 [] 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 [] c (* (nb_lam [na1:T1]...[nan:Tan]c) where c is not an abstraction * gives n (casts are ignored) *) @@ -1007,7 +1000,7 @@ let rec eta_reduce_head sigma 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 @@ -1055,6 +1048,14 @@ 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 Environ.lookup_named_val id ctxt; true with Not_found -> false +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 = let compact l decl = match decl, l with @@ -1098,10 +1099,10 @@ let global_vars_set env sigma constr = let global_vars env sigma ids = Id.Set.elements (global_vars_set env sigma ids) let global_vars_set_of_decl env sigma = function - | NamedDecl.LocalAssum (_,t) -> global_vars_set env sigma (EConstr.of_constr t) + | NamedDecl.LocalAssum (_,t) -> global_vars_set env sigma t | NamedDecl.LocalDef (_,c,t) -> - Id.Set.union (global_vars_set env sigma (EConstr.of_constr t)) - (global_vars_set env sigma (EConstr.of_constr c)) + Id.Set.union (global_vars_set env sigma t) + (global_vars_set env sigma c) let dependency_closure env sigma sign hyps = if Id.Set.is_empty hyps then [] else @@ -1155,6 +1156,7 @@ let context_chop 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), diff --git a/engine/termops.mli b/engine/termops.mli index 196962354..841de34b1 100644 --- a/engine/termops.mli +++ b/engine/termops.mli @@ -34,7 +34,7 @@ val push_rel_assum : Name.t * types -> 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.constr option * Constr.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. *) @@ -45,16 +45,16 @@ 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_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 -> Context.Named.t -> types +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 -> Context.Named.t -> constr +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 *) @@ -64,12 +64,12 @@ val it_mkLambda_or_LetIn_from_no_LetIn : Constr.constr -> Context.Rel.t -> Const val map_constr_with_binders_left_to_right : Evd.evar_map -> - (Context.Rel.Declaration.t -> 'a -> 'a) -> + (rel_declaration -> 'a -> 'a) -> ('a -> constr -> constr) -> 'a -> constr -> constr val map_constr_with_full_binders : Evd.evar_map -> - (Context.Rel.Declaration.t -> 'a -> 'a) -> + (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 @@ -105,7 +105,7 @@ 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 -> Evd.evar_map -> - Id.t -> Context.Named.Declaration.t -> bool + Id.t -> named_declaration -> bool (** As {!occur_var} but assume the identifier not to be a section variable *) val local_occur_var : Evd.evar_map -> Id.t -> constr -> bool @@ -117,7 +117,7 @@ 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 -> Context.Named.Declaration.t -> 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 *) @@ -182,11 +182,11 @@ 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.constr) Evar.Map.t -val filtering : Evd.evar_map -> Context.Rel.t -> Reduction.conv_pb -> Constr.constr -> 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 : Evd.evar_map -> constr -> int * Context.Rel.t * constr -val align_prod_letin : Evd.evar_map -> 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) *) @@ -215,8 +215,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 @@ -228,15 +228,15 @@ 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.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.constr list -> Context.Rel.t -> Context.Rel.t val smash_rel_context : Context.Rel.t -> Context.Rel.t (** expand lets in context *) @@ -244,23 +244,26 @@ val smash_rel_context : Context.Rel.t -> Context.Rel.t (** expand lets in contex val map_rel_context_in_env : (env -> Constr.constr -> Constr.constr) -> env -> Context.Rel.t -> Context.Rel.t val map_rel_context_with_binders : - (int -> Constr.constr -> 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.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 -> 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 -> Context.Named.Declaration.t -> 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 puniverses * constr option (** Gives an ordered list of hypotheses, closed by dependencies, containing a given set *) -val dependency_closure : env -> Evd.evar_map -> 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 diff --git a/interp/constrintern.ml b/interp/constrintern.ml index e5dd6a6ec..41e2e4e6f 100644 --- a/interp/constrintern.ml +++ b/interp/constrintern.ml @@ -2083,6 +2083,7 @@ let intern_context global_level env impl_env binders = user_err ~loc ~hdr:"internalize" (explain_internalization_error e) let interp_rawcontext_evars env evdref k bl = + let open EConstr in let (env, par, _, impls) = List.fold_left (fun (env,params,n,impls) (na, k, b, t) -> @@ -2093,7 +2094,6 @@ let interp_rawcontext_evars env evdref k bl = let t = understand_tcc_evars env evdref ~expected_type:IsType t' in match b with None -> - let t = EConstr.Unsafe.to_constr t in let d = LocalAssum (na,t) in let impls = if k == Implicit then @@ -2104,8 +2104,6 @@ let interp_rawcontext_evars env evdref k bl = (push_rel d env, d::params, succ n, impls) | Some b -> let c = understand_tcc_evars env evdref ~expected_type:(OfType t) b in - let t = EConstr.Unsafe.to_constr t in - let c = EConstr.Unsafe.to_constr c in let d = LocalDef (na, c, t) in (push_rel d env, d::params, n, impls)) (env,[],k+1,[]) (List.rev bl) diff --git a/interp/constrintern.mli b/interp/constrintern.mli index 2f6795ed4..ae7f511f4 100644 --- a/interp/constrintern.mli +++ b/interp/constrintern.mli @@ -160,7 +160,7 @@ val interp_binder_evars : env -> evar_map ref -> Name.t -> constr_expr -> EConst val interp_context_evars : ?global_level:bool -> ?impl_env:internalization_env -> ?shift:int -> env -> evar_map ref -> local_binder list -> - internalization_env * ((env * Context.Rel.t) * Impargs.manual_implicits) + internalization_env * ((env * EConstr.rel_context) * Impargs.manual_implicits) (* val interp_context_gen : (env -> glob_constr -> unsafe_type_judgment Evd.in_evar_universe_context) -> *) (* (env -> Evarutil.type_constraint -> glob_constr -> unsafe_judgment Evd.in_evar_universe_context) -> *) @@ -177,7 +177,7 @@ val interp_context_evars : val locate_reference : Libnames.qualid -> Globnames.global_reference val is_global : Id.t -> bool -val construct_reference : Context.Named.t -> Id.t -> constr +val construct_reference : ('c, 't) Context.Named.pt -> Id.t -> constr val global_reference : Id.t -> constr val global_reference_in_absolute_module : DirPath.t -> Id.t -> constr diff --git a/ltac/rewrite.ml b/ltac/rewrite.ml index c942e8e92..b45ead217 100644 --- a/ltac/rewrite.ml +++ b/ltac/rewrite.ml @@ -41,22 +41,6 @@ open Context.Named.Declaration module NamedDecl = Context.Named.Declaration module RelDecl = Context.Rel.Declaration -let local_assum (na, t) = - let inj = EConstr.Unsafe.to_constr in - RelDecl.LocalAssum (na, inj t) - -let local_def (na, b, t) = - let inj = EConstr.Unsafe.to_constr in - RelDecl.LocalDef (na, inj b, inj t) - -let nlocal_assum (na, t) = - let inj = EConstr.Unsafe.to_constr in - NamedDecl.LocalAssum (na, inj t) - -let nlocal_def (na, b, t) = - let inj = EConstr.Unsafe.to_constr in - NamedDecl.LocalDef (na, inj b, inj t) - (** Typeclass-based generalized rewriting. *) (** Constants used by the tactic. *) @@ -251,7 +235,7 @@ end) = struct evars, mkProd(na, ty, b), newarg, (ty, Some relty) :: cstrs else let (evars, b, arg, cstrs) = - aux (Environ.push_rel (local_assum (na, ty)) env) evars b cstrs + aux (push_rel (LocalAssum (na, ty)) env) evars b cstrs in let ty = Reductionops.nf_betaiota (goalevars evars) ty in let pred = mkLambda (na, ty, b) in @@ -350,7 +334,7 @@ end) = struct let evars, rb = aux evars env b' (pred n) in app_poly env evars pointwise_relation [| ty; b'; rb |] else - let evars, rb = aux evars (Environ.push_rel (local_assum (na, ty)) env) b (pred n) in + let evars, rb = aux evars (push_rel (LocalAssum (na, ty)) env) b (pred n) in app_poly env evars forall_relation [| ty; mkLambda (na, ty, b); mkLambda (na, ty, rb) |] | _ -> raise Not_found @@ -379,7 +363,7 @@ end) = struct else (try let params, args = Array.chop (Array.length args - 2) args in - let env' = Environ.push_rel_context rels env in + let env' = push_rel_context rels env in let sigma = Sigma.Unsafe.of_evar_map sigma in let Sigma ((evar, _), evars, _) = Evarutil.new_type_evar env' sigma Evd.univ_flexible in let evars = Sigma.to_evar_map evars in @@ -546,7 +530,7 @@ let decompose_applied_relation env sigma (c,l) = | Some c -> c | None -> let ctx,t' = Reductionops.splay_prod env sigma ctype in (* Search for underlying eq *) - match find_rel (it_mkProd_or_LetIn t' (List.map (fun (n,t) -> local_assum (n, t)) ctx)) with + match find_rel (it_mkProd_or_LetIn t' (List.map (fun (n,t) -> LocalAssum (n, t)) ctx)) with | Some c -> c | None -> error "Cannot find an homogeneous relation to rewrite." @@ -826,8 +810,8 @@ let resolve_morphism env avoid oldt m ?(fnewt=fun x -> x) args args' (b,cstr) ev if b then PropGlobal.do_subrelation, PropGlobal.apply_subrelation else TypeGlobal.do_subrelation, TypeGlobal.apply_subrelation in - Environ.push_named - (nlocal_def (Id.of_string "do_subrelation", + EConstr.push_named + (LocalDef (Id.of_string "do_subrelation", snd (app_poly_sort b env evars dosub [||]), snd (app_poly_nocheck env evars appsub [||]))) env @@ -947,7 +931,7 @@ let fold_match ?(force=false) env sigma c = let dep, pred, exists, (sk,eff) = let env', ctx, body = let ctx, pred = decompose_lam_assum sigma p in - let env' = Environ.push_rel_context ctx env in + let env' = push_rel_context ctx env in env', ctx, pred in let sortp = Retyping.get_sort_family_of env' sigma body in @@ -1174,7 +1158,7 @@ let subterm all flags (s : 'a pure_strategy) : 'a pure_strategy = | Lambda (n, t, b) when flags.under_lambdas -> let n' = name_app (fun id -> Tactics.fresh_id_in_env unfresh id env) n in let open Context.Rel.Declaration in - let env' = Environ.push_rel (local_assum (n', t)) env in + let env' = EConstr.push_rel (LocalAssum (n', t)) env in let bty = Retyping.get_type_of env' (goalevars evars) b in let unlift = if prop then PropGlobal.unlift_cstr else TypeGlobal.unlift_cstr in let state, b' = s.strategy { state ; env = env' ; unfresh ; @@ -1563,15 +1547,14 @@ let assert_replacing id newt tac = let concl = Proofview.Goal.concl gl in let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in - let ctx = Environ.named_context env in + let ctx = named_context env in let after, before = List.split_when (NamedDecl.get_id %> Id.equal id) ctx in let nc = match before with | [] -> assert false - | d :: rem -> insert_dependent env sigma (nlocal_assum (NamedDecl.get_id d, newt)) [] after @ rem + | d :: rem -> insert_dependent env sigma (LocalAssum (NamedDecl.get_id d, newt)) [] after @ rem in let env' = Environ.reset_with_named_context (val_of_named_context nc) env in Refine.refine ~unsafe:false { run = begin fun sigma -> - let open EConstr in let Sigma (ev, sigma, p) = Evarutil.new_evar env' sigma concl in let Sigma (ev', sigma, q) = Evarutil.new_evar env sigma newt in let map d = @@ -1612,7 +1595,7 @@ let cl_rewrite_clause_newtac ?abs ?origsigma ~progress strat clause = tclTHENFIRST (assert_replacing id newt tac) (beta_hyp id) | Some id, None -> Proofview.Unsafe.tclEVARS undef <*> - convert_hyp_no_check (nlocal_assum (id, newt)) <*> + convert_hyp_no_check (LocalAssum (id, newt)) <*> beta_hyp id | None, Some p -> Proofview.Unsafe.tclEVARS undef <*> @@ -1640,10 +1623,10 @@ let cl_rewrite_clause_newtac ?abs ?origsigma ~progress strat clause = | None -> env | Some id -> (** Only consider variables not depending on [id] *) - let ctx = Environ.named_context env in + let ctx = named_context env in let filter decl = not (occur_var_in_decl env sigma id decl) in let nctx = List.filter filter ctx in - Environ.reset_with_named_context (Environ.val_of_named_context nctx) env + Environ.reset_with_named_context (val_of_named_context nctx) env in try let res = diff --git a/ltac/rewrite.mli b/ltac/rewrite.mli index 378359d51..10fec2032 100644 --- a/ltac/rewrite.mli +++ b/ltac/rewrite.mli @@ -8,8 +8,8 @@ open Names open Constr -open EConstr open Environ +open EConstr open Constrexpr open Tacexpr open Misctypes @@ -75,7 +75,7 @@ val cl_rewrite_clause : bool -> Locus.occurrences -> Id.t option -> unit Proofview.tactic val is_applied_rewrite_relation : - env -> evar_map -> Context.Rel.t -> constr -> types option + env -> evar_map -> rel_context -> constr -> types option val declare_relation : ?binders:local_binder list -> constr_expr -> constr_expr -> Id.t -> diff --git a/ltac/tactic_matching.ml b/ltac/tactic_matching.ml index ac64f0bba..5b5cd06cc 100644 --- a/ltac/tactic_matching.ml +++ b/ltac/tactic_matching.ml @@ -284,7 +284,7 @@ module PatternMatching (E:StaticEnvironment) = struct pick hyps >>= fun decl -> let id = NamedDecl.get_id decl in let refresh = is_local_def decl in - pattern_match_term refresh pat (EConstr.of_constr (NamedDecl.get_type decl)) () <*> + pattern_match_term refresh pat (NamedDecl.get_type decl) () <*> put_terms (id_map_try_add_name hypname (EConstr.mkVar id) empty_term_subst) <*> return id @@ -295,8 +295,6 @@ module PatternMatching (E:StaticEnvironment) = struct let hyp_match_body_and_type hypname bodypat typepat hyps = pick hyps >>= function | LocalDef (id,body,hyp) -> - let body = EConstr.of_constr body in - let hyp = EConstr.of_constr hyp in pattern_match_term false bodypat body () <*> pattern_match_term true typepat hyp () <*> put_terms (id_map_try_add_name hypname (EConstr.mkVar id) empty_term_subst) <*> diff --git a/ltac/tactic_matching.mli b/ltac/tactic_matching.mli index d6f67d6c7..300b546f1 100644 --- a/ltac/tactic_matching.mli +++ b/ltac/tactic_matching.mli @@ -43,7 +43,7 @@ val match_term : val match_goal: Environ.env -> Evd.evar_map -> - Context.Named.t -> + EConstr.named_context -> EConstr.constr -> (Tacexpr.binding_bound_vars * Pattern.constr_pattern, Tacexpr.glob_tactic_expr) Tacexpr.match_rule list -> Tacexpr.glob_tactic_expr t Proofview.tactic diff --git a/plugins/cc/cctac.ml b/plugins/cc/cctac.ml index 53c450116..5d894c677 100644 --- a/plugins/cc/cctac.ml +++ b/plugins/cc/cctac.ml @@ -28,10 +28,6 @@ open Proofview.Notations module RelDecl = Context.Rel.Declaration module NamedDecl = Context.Named.Declaration -let local_assum (na, t) = - let inj = EConstr.Unsafe.to_constr in - RelDecl.LocalAssum (na, inj t) - let reference dir s = lazy (Coqlib.gen_reference "CC" dir s) let _f_equal = reference ["Init";"Logic"] "f_equal" @@ -160,7 +156,7 @@ let rec quantified_atom_of_constr env sigma nrels term = let patts=patterns_of_constr env sigma nrels atom in `Nrule patts else - quantified_atom_of_constr (Environ.push_rel (local_assum (id,atom)) env) sigma (succ nrels) ff + quantified_atom_of_constr (EConstr.push_rel (RelDecl.LocalAssum (id,atom)) env) sigma (succ nrels) ff | _ -> let patts=patterns_of_constr env sigma nrels term in `Rule patts @@ -175,7 +171,7 @@ let litteral_of_constr env sigma term= else begin try - quantified_atom_of_constr (Environ.push_rel (local_assum (id,atom)) env) sigma 1 ff + quantified_atom_of_constr (EConstr.push_rel (RelDecl.LocalAssum (id,atom)) env) sigma 1 ff with Not_found -> `Other (decompose_term env sigma term) end diff --git a/plugins/decl_mode/decl_proof_instr.ml b/plugins/decl_mode/decl_proof_instr.ml index da971fffb..adc4ad8a3 100644 --- a/plugins/decl_mode/decl_proof_instr.ml +++ b/plugins/decl_mode/decl_proof_instr.ml @@ -49,7 +49,7 @@ let clear ids { it = goal; sigma } = user_err (str "Cannot clear " ++ pr_id id) in let sigma = !evdref in - let (gl,ev,sigma) = Goal.V82.mk_goal sigma hyps (EConstr.of_constr concl) (Goal.V82.extra sigma goal) in + let (gl,ev,sigma) = Goal.V82.mk_goal sigma hyps concl (Goal.V82.extra sigma goal) in let sigma = Goal.V82.partial_solution_to sigma goal gl ev in { it = [gl]; sigma } @@ -638,7 +638,7 @@ let assume_tac hyps gls = tclTHEN (push_intro_tac (fun id -> - Proofview.V82.of_tactic (convert_hyp (LocalAssum (id,st.st_it)))) st.st_label)) + Proofview.V82.of_tactic (convert_hyp (LocalAssum (id,EConstr.of_constr st.st_it)))) st.st_label)) hyps tclIDTAC gls let assume_hyps_or_theses hyps gls = @@ -648,7 +648,7 @@ let assume_hyps_or_theses hyps gls = tclTHEN (push_intro_tac (fun id -> - Proofview.V82.of_tactic (convert_hyp (LocalAssum (id,c)))) nam) + Proofview.V82.of_tactic (convert_hyp (LocalAssum (id,EConstr.of_constr c)))) nam) | Hprop {st_label=nam;st_it=Thesis (tk)} -> tclTHEN (push_intro_tac @@ -660,7 +660,7 @@ let assume_st hyps gls = (fun st -> tclTHEN (push_intro_tac - (fun id -> Proofview.V82.of_tactic (convert_hyp (LocalAssum (id,st.st_it)))) st.st_label)) + (fun id -> Proofview.V82.of_tactic (convert_hyp (LocalAssum (id,EConstr.of_constr st.st_it)))) st.st_label)) hyps tclIDTAC gls let assume_st_letin hyps gls = @@ -669,7 +669,7 @@ let assume_st_letin hyps gls = tclTHEN (push_intro_tac (fun id -> - Proofview.V82.of_tactic (convert_hyp (LocalDef (id, fst st.st_it, snd st.st_it)))) st.st_label)) + Proofview.V82.of_tactic (convert_hyp (LocalDef (id, EConstr.of_constr (fst st.st_it), EConstr.of_constr (snd st.st_it))))) st.st_label)) hyps tclIDTAC gls (* suffices *) @@ -763,7 +763,7 @@ let rec consider_match may_intro introduced available expected gls = error "Not enough sub-hypotheses to match statements." (* should tell which ones *) | id::rest_ids,(Hvar st | Hprop st)::rest -> - tclIFTHENELSE (Proofview.V82.of_tactic (convert_hyp (LocalAssum (id,st.st_it)))) + tclIFTHENELSE (Proofview.V82.of_tactic (convert_hyp (LocalAssum (id,EConstr.of_constr st.st_it)))) begin match st.st_label with Anonymous -> @@ -834,13 +834,13 @@ let define_tac id args body gls = (* tactics for reconsider *) let cast_tac id_or_thesis typ gls = + let typ = EConstr.of_constr typ in match id_or_thesis with | This id -> Proofview.V82.of_tactic (id |> pf_get_hyp gls |> NamedDecl.set_id id |> NamedDecl.set_type typ |> convert_hyp) gls | Thesis (For _ ) -> error "\"thesis for ...\" is not applicable here." | Thesis Plain -> - let typ = EConstr.of_constr typ in Proofview.V82.of_tactic (convert_concl typ DEFAULTcast) gls (* per cases *) @@ -1290,6 +1290,7 @@ let rec execute_cases fix_name per_info tacnext args objs nhrec tree gls = let f_ids typ = let sign = (prod_assum (Term.prod_applist typ params)) in + let sign = List.map (fun d -> Termops.map_rel_decl EConstr.of_constr d) sign in find_intro_names sign gls in let constr_args_ids = Array.map f_ids gen_arities in let case_term = diff --git a/plugins/extraction/extraction.ml b/plugins/extraction/extraction.ml index 7b7e746f2..8744eacd3 100644 --- a/plugins/extraction/extraction.ml +++ b/plugins/extraction/extraction.ml @@ -858,7 +858,7 @@ and extract_fix env mle i (fi,ti,ci as recd) mlt = let decomp_lams_eta_n n m env c t = let rels = fst (splay_prod_n env none n (EConstr.of_constr t)) in - let rels = List.map (fun (LocalAssum (id,c) | LocalDef (id,_,c)) -> (id,c)) rels in + let rels = List.map (fun (LocalAssum (id,c) | LocalDef (id,_,c)) -> (id,EConstr.Unsafe.to_constr c)) rels in let rels',c = decompose_lam c in let d = n - m in (* we'd better keep rels' as long as possible. *) diff --git a/plugins/firstorder/instances.ml b/plugins/firstorder/instances.ml index ef8172de4..9dc2a51a6 100644 --- a/plugins/firstorder/instances.ml +++ b/plugins/firstorder/instances.ml @@ -117,10 +117,9 @@ let mk_open_instance id idc gl m t= let nid=(fresh_id avoid var_id gl) in let evmap = Sigma.Unsafe.of_evar_map evmap in let Sigma ((c, _), evmap, _) = Evarutil.new_type_evar env evmap Evd.univ_flexible in - let c = EConstr.Unsafe.to_constr c in let evmap = Sigma.to_evar_map evmap in let decl = LocalAssum (Name nid, c) in - aux (n-1) (nid::avoid) (Environ.push_rel decl env) evmap (decl::decls) in + aux (n-1) (nid::avoid) (EConstr.push_rel decl env) evmap (decl::decls) in let evmap, decls = aux m [] env evmap [] in evmap, decls, revt diff --git a/plugins/firstorder/rules.ml b/plugins/firstorder/rules.ml index 36bd91ab6..a60fd4d8f 100644 --- a/plugins/firstorder/rules.ml +++ b/plugins/firstorder/rules.ml @@ -42,7 +42,7 @@ let wrap n b continue seq gls= List.exists (occur_var_in_decl env (project gls) id) ctx then (aux (i-1) q (nd::ctx)) else - add_formula Hyp (VarRef id) (NamedDecl.get_type nd) (aux (i-1) q (nd::ctx)) gls in + add_formula Hyp (VarRef id) (EConstr.Unsafe.to_constr (NamedDecl.get_type nd)) (aux (i-1) q (nd::ctx)) gls in let seq1=aux n nc [] in let seq2=if b then add_formula Concl dummy_id (EConstr.Unsafe.to_constr (pf_concl gls)) seq1 gls else seq1 in diff --git a/plugins/fourier/fourierR.ml b/plugins/fourier/fourierR.ml index ec73fccb5..e11cbc279 100644 --- a/plugins/fourier/fourierR.ml +++ b/plugins/fourier/fourierR.ml @@ -190,6 +190,8 @@ type hineq={hname:constr; (* le nom de l'hypothèse *) exception NoIneq let ineq1_of_constr (h,t) = + let h = EConstr.Unsafe.to_constr h in + let t = EConstr.Unsafe.to_constr t in match (kind_of_term t) with | App (f,args) -> (match kind_of_term f with @@ -504,7 +506,7 @@ let rec fourier () = |_-> raise GoalDone with GoalDone -> (* les hypothèses *) - let hyps = List.map (fun (h,t)-> (mkVar h,t)) + let hyps = List.map (fun (h,t)-> (EConstr.mkVar h,t)) (list_of_sign (Proofview.Goal.hyps gl)) in let lineq =ref [] in List.iter (fun h -> try (lineq:=(ineq1_of_constr h)@(!lineq)) diff --git a/plugins/funind/functional_principles_proofs.ml b/plugins/funind/functional_principles_proofs.ml index 91b17b9a4..bc64b079c 100644 --- a/plugins/funind/functional_principles_proofs.ml +++ b/plugins/funind/functional_principles_proofs.ml @@ -236,7 +236,7 @@ let nf_betaiotazeta = (* Reductionops.local_strong Reductionops.whd_betaiotazeta -let change_eq env sigma hyp_id (context:Context.Rel.t) x t end_of_type = +let change_eq env sigma hyp_id (context:rel_context) x t end_of_type = let nochange ?t' msg = begin observe (str ("Not treating ( "^msg^" )") ++ pr_leconstr t ++ str " " ++ match t' with None -> str "" | Some t -> Printer.pr_leconstr t ); @@ -315,7 +315,7 @@ let change_eq env sigma hyp_id (context:Context.Rel.t) x t end_of_type = try let witness = Int.Map.find i sub in if is_local_def decl then anomaly (Pp.str "can not redefine a rel!"); - (pop end_of_type,ctxt_size,mkLetIn (RelDecl.get_name decl, witness, EConstr.of_constr (RelDecl.get_type decl), witness_fun)) + (pop end_of_type,ctxt_size,mkLetIn (RelDecl.get_name decl, witness, RelDecl.get_type decl, witness_fun)) with Not_found -> (mkProd_or_LetIn decl end_of_type, ctxt_size + 1, mkLambda_or_LetIn decl witness_fun) ) @@ -544,7 +544,7 @@ let clean_hyp_with_heq ptes_infos eq_hyps hyp_id env sigma = (scan_type new_context new_t') with Failure "NoChange" -> (* Last thing todo : push the rel in the context and continue *) - scan_type (local_assum (x,t_x) :: context) t' + scan_type (LocalAssum (x,t_x) :: context) t' end end else @@ -933,6 +933,7 @@ let generalize_non_dep hyp g = let to_revert,_ = let open Context.Named.Declaration in Environ.fold_named_context_reverse (fun (clear,keep) decl -> + let decl = map_named_decl EConstr.of_constr decl in let hyp = get_id decl in if Id.List.mem hyp hyps || List.exists (Termops.occur_var_in_decl env (project g) hyp) keep diff --git a/plugins/funind/functional_principles_types.ml b/plugins/funind/functional_principles_types.ml index d0d44b34b..e845db3bc 100644 --- a/plugins/funind/functional_principles_types.ml +++ b/plugins/funind/functional_principles_types.ml @@ -33,9 +33,9 @@ let compute_new_princ_type_from_rel rel_to_fun sorts princ_type = let princ_type = EConstr.of_constr princ_type in let princ_type_info = compute_elim_sig Evd.empty princ_type (** FIXME *) in let env = Global.env () in - let env_with_params = Environ.push_rel_context princ_type_info.params env in + let env_with_params = EConstr.push_rel_context princ_type_info.params env in let tbl = Hashtbl.create 792 in - let rec change_predicates_names (avoid:Id.t list) (predicates:Context.Rel.t) : Context.Rel.t = + let rec change_predicates_names (avoid:Id.t list) (predicates:EConstr.rel_context) : EConstr.rel_context = match predicates with | [] -> [] | decl :: predicates -> @@ -56,7 +56,7 @@ let compute_new_princ_type_from_rel rel_to_fun sorts princ_type = (* observe (str "princ_infos : " ++ pr_elim_scheme princ_type_info); *) let change_predicate_sort i decl = let new_sort = sorts.(i) in - let args,_ = decompose_prod (RelDecl.get_type decl) in + let args,_ = decompose_prod (EConstr.Unsafe.to_constr (RelDecl.get_type decl)) in let real_args = if princ_type_info.indarg_in_concl then List.tl args @@ -87,17 +87,19 @@ let compute_new_princ_type_from_rel rel_to_fun sorts princ_type = | _ -> false in let pre_princ = + let open EConstr in it_mkProd_or_LetIn (it_mkProd_or_LetIn (Option.fold_right mkProd_or_LetIn princ_type_info.indarg - (EConstr.Unsafe.to_constr princ_type_info.concl) + princ_type_info.concl ) princ_type_info.args ) princ_type_info.branches in + let pre_princ = EConstr.Unsafe.to_constr pre_princ in let pre_princ = substl (List.map mkVar ptes_vars) pre_princ in let is_dom c = match kind_of_term c with @@ -240,7 +242,7 @@ let compute_new_princ_type_from_rel rel_to_fun sorts princ_type = | Context.Named.Declaration.LocalDef (id,t,b) -> LocalDef (Name (Hashtbl.find tbl id), t, b)) new_predicates) ) - princ_type_info.params + (List.map (fun d -> Termops.map_rel_decl EConstr.Unsafe.to_constr d) princ_type_info.params) @@ -251,7 +253,7 @@ let change_property_sort evd toSort princ princName = let change_sort_in_predicate decl = LocalAssum (get_name decl, - let args,ty = decompose_prod (get_type decl) in + let args,ty = decompose_prod (EConstr.Unsafe.to_constr (get_type decl)) in let s = destSort ty in Global.add_constraints (Univ.enforce_leq (univ_of_sort toSort) (univ_of_sort s) Univ.Constraint.empty); Term.compose_prod args (mkSort toSort) @@ -270,7 +272,7 @@ let change_property_sort evd toSort princ princName = (it_mkLambda_or_LetIn init (List.map change_sort_in_predicate princ_info.predicates) ) - princ_info.params + (List.map (fun d -> Termops.map_rel_decl EConstr.Unsafe.to_constr d) princ_info.params) let build_functional_principle (evd:Evd.evar_map ref) interactive_proof old_princ_type sorts funs i proof_tac hook = (* First we get the type of the old graph principle *) diff --git a/plugins/funind/indfun.ml b/plugins/funind/indfun.ml index 1cde4420e..a7489fb7b 100644 --- a/plugins/funind/indfun.ml +++ b/plugins/funind/indfun.ml @@ -19,7 +19,7 @@ let is_rec_info sigma scheme_info = let test_branche min acc decl = acc || ( let new_branche = - it_mkProd_or_LetIn mkProp (fst (decompose_prod_assum sigma (EConstr.of_constr (RelDecl.get_type decl)))) in + it_mkProd_or_LetIn mkProp (fst (decompose_prod_assum sigma (RelDecl.get_type decl))) in let free_rels_in_br = Termops.free_rels sigma new_branche in let max = min + scheme_info.Tactics.npredicates in Int.Set.exists (fun i -> i >= min && i< max) free_rels_in_br diff --git a/plugins/funind/invfun.ml b/plugins/funind/invfun.ml index dcec2cb74..8f1420940 100644 --- a/plugins/funind/invfun.ml +++ b/plugins/funind/invfun.ml @@ -26,12 +26,6 @@ open Context.Rel.Declaration module RelDecl = Context.Rel.Declaration -let local_assum (na, t) = - RelDecl.LocalAssum (na, EConstr.Unsafe.to_constr t) - -let local_def (na, b, t) = - RelDecl.LocalDef (na, EConstr.Unsafe.to_constr b, EConstr.Unsafe.to_constr t) - (* Some pretty printing function for debugging purpose *) let pr_binding prc = @@ -147,7 +141,7 @@ let generate_type evd g_to_f f graph i = let fun_ctxt,res_type = match ctxt with | [] | [_] -> anomaly (Pp.str "Not a valid context") - | decl :: fun_ctxt -> fun_ctxt, EConstr.of_constr (RelDecl.get_type decl) + | decl :: fun_ctxt -> fun_ctxt, RelDecl.get_type decl in let rec args_from_decl i accu = function | [] -> accu @@ -187,12 +181,12 @@ let generate_type evd g_to_f f graph i = \[\forall (x_1:t_1)\ldots(x_n:t_n), let fv := f x_1\ldots x_n in, forall res, \] i*) let pre_ctxt = - local_assum (Name res_id, lift 1 res_type) :: local_def (Name fv_id, mkApp (f,args_as_rels), res_type) :: fun_ctxt + LocalAssum (Name res_id, lift 1 res_type) :: LocalDef (Name fv_id, mkApp (f,args_as_rels), res_type) :: fun_ctxt in (*i and we can return the solution depending on which lemma type we are defining i*) if g_to_f - then local_assum (Anonymous,graph_applied)::pre_ctxt,(lift 1 res_eq_f_of_args),graph - else local_assum (Anonymous,res_eq_f_of_args)::pre_ctxt,(lift 1 graph_applied),graph + then LocalAssum (Anonymous,graph_applied)::pre_ctxt,(lift 1 res_eq_f_of_args),graph + else LocalAssum (Anonymous,res_eq_f_of_args)::pre_ctxt,(lift 1 graph_applied),graph (* @@ -280,7 +274,7 @@ let prove_fun_correct evd functional_induction funs_constr graphs_constr schemes (fun decl -> List.map (fun id -> Loc.ghost, IntroNaming (IntroIdentifier id)) - (generate_fresh_id (Id.of_string "y") ids (List.length (fst (decompose_prod_assum evd (EConstr.of_constr (RelDecl.get_type decl)))))) + (generate_fresh_id (Id.of_string "y") ids (List.length (fst (decompose_prod_assum evd (RelDecl.get_type decl))))) ) branches in @@ -477,7 +471,7 @@ let generalize_dependent_of x hyp g = tclMAP (function | LocalAssum (id,t) when not (Id.equal id hyp) && - (Termops.occur_var (pf_env g) (project g) x (EConstr.of_constr t)) -> tclTHEN (Proofview.V82.of_tactic (Tactics.generalize [mkVar id])) (thin [id]) + (Termops.occur_var (pf_env g) (project g) x t) -> tclTHEN (Proofview.V82.of_tactic (Tactics.generalize [mkVar id])) (thin [id]) | _ -> tclIDTAC ) (pf_hyps g) @@ -695,7 +689,7 @@ let prove_fun_complete funcs graphs schemes lemmas_types_infos i : tactic = (fun decl -> List.map (fun id -> id) - (generate_fresh_id (Id.of_string "y") ids (nb_prod (project g) (EConstr.of_constr (RelDecl.get_type decl)))) + (generate_fresh_id (Id.of_string "y") ids (nb_prod (project g) (RelDecl.get_type decl))) ) branches in diff --git a/plugins/funind/merge.ml b/plugins/funind/merge.ml index 2840193a9..691385fad 100644 --- a/plugins/funind/merge.ml +++ b/plugins/funind/merge.ml @@ -976,7 +976,7 @@ let funify_branches relinfo nfuns branch = | Rel i -> let reali = i-shift in (reali>=0 && reali false in (* FIXME: *) - LocalDef (Anonymous,mkProp,mkProp) + LocalDef (Anonymous,EConstr.mkProp,EConstr.mkProp) let relprinctype_to_funprinctype relprinctype nfuns = diff --git a/plugins/funind/recdef.ml b/plugins/funind/recdef.ml index 56c6ab054..f5ea32878 100644 --- a/plugins/funind/recdef.ml +++ b/plugins/funind/recdef.ml @@ -693,7 +693,7 @@ let mkDestructEq : (fun decl -> let open Context.Named.Declaration in let id = get_id decl in - if Id.List.mem id not_on_hyp || not (Termops.occur_term (project g) expr (EConstr.of_constr (get_type decl))) + if Id.List.mem id not_on_hyp || not (Termops.occur_term (project g) expr (get_type decl)) then None else Some id) hyps in let to_revert_constr = List.rev_map mkVar to_revert in let type_of_expr = pf_unsafe_type_of g expr in diff --git a/plugins/omega/coq_omega.ml b/plugins/omega/coq_omega.ml index 9e0d591b6..8c2f0f53f 100644 --- a/plugins/omega/coq_omega.ml +++ b/plugins/omega/coq_omega.ml @@ -1706,10 +1706,6 @@ let onClearedName2 id tac = Tacticals.New.tclTHENLIST [ introduction id1; introduction id2; tac id1 id2 ] end }) -let local_assum (na, t) = - let inj = EConstr.Unsafe.to_constr in - LocalAssum (na, inj t) - let rec is_Prop sigma c = match EConstr.kind sigma c with | Sort (Prop Null) -> true | Cast (c,_,_) -> is_Prop sigma c @@ -1725,24 +1721,24 @@ let destructure_hyps = | decl::lit -> let i = NamedDecl.get_id decl in Proofview.tclEVARMAP >>= fun sigma -> - begin try match destructurate_prop sigma (EConstr.of_constr (NamedDecl.get_type decl)) with + begin try match destructurate_prop sigma (NamedDecl.get_type decl) with | Kapp(False,[]) -> elim_id i | Kapp((Zle|Zge|Zgt|Zlt|Zne),[t1;t2]) -> loop lit | Kapp(Or,[t1;t2]) -> (Tacticals.New.tclTHENS (elim_id i) - [ onClearedName i (fun i -> (loop (local_assum (i,t1)::lit))); - onClearedName i (fun i -> (loop (local_assum (i,t2)::lit))) ]) + [ onClearedName i (fun i -> (loop (LocalAssum (i,t1)::lit))); + onClearedName i (fun i -> (loop (LocalAssum (i,t2)::lit))) ]) | Kapp(And,[t1;t2]) -> Tacticals.New.tclTHEN (elim_id i) (onClearedName2 i (fun i1 i2 -> - loop (local_assum (i1,t1) :: local_assum (i2,t2) :: lit))) + loop (LocalAssum (i1,t1) :: LocalAssum (i2,t2) :: lit))) | Kapp(Iff,[t1;t2]) -> Tacticals.New.tclTHEN (elim_id i) (onClearedName2 i (fun i1 i2 -> - loop (local_assum (i1,mkArrow t1 t2) :: local_assum (i2,mkArrow t2 t1) :: lit))) + loop (LocalAssum (i1,mkArrow t1 t2) :: LocalAssum (i2,mkArrow t2 t1) :: lit))) | Kimp(t1,t2) -> (* t1 and t2 might be in Type rather than Prop. For t1, the decidability check will ensure being Prop. *) @@ -1753,7 +1749,7 @@ let destructure_hyps = (generalize_tac [mkApp (Lazy.force coq_imp_simp, [| t1; t2; d1; mkVar i|])]); (onClearedName i (fun i -> - (loop (local_assum (i,mk_or (mk_not t1) t2) :: lit)))) + (loop (LocalAssum (i,mk_or (mk_not t1) t2) :: lit)))) ] else loop lit @@ -1764,7 +1760,7 @@ let destructure_hyps = (generalize_tac [mkApp (Lazy.force coq_not_or,[| t1; t2; mkVar i |])]); (onClearedName i (fun i -> - (loop (local_assum (i,mk_and (mk_not t1) (mk_not t2)) :: lit)))) + (loop (LocalAssum (i,mk_and (mk_not t1) (mk_not t2)) :: lit)))) ] | Kapp(And,[t1;t2]) -> let d1 = decidability t1 in @@ -1773,7 +1769,7 @@ let destructure_hyps = [mkApp (Lazy.force coq_not_and, [| t1; t2; d1; mkVar i |])]); (onClearedName i (fun i -> - (loop (local_assum (i,mk_or (mk_not t1) (mk_not t2)) :: lit)))) + (loop (LocalAssum (i,mk_or (mk_not t1) (mk_not t2)) :: lit)))) ] | Kapp(Iff,[t1;t2]) -> let d1 = decidability t1 in @@ -1783,7 +1779,7 @@ let destructure_hyps = [mkApp (Lazy.force coq_not_iff, [| t1; t2; d1; d2; mkVar i |])]); (onClearedName i (fun i -> - (loop (local_assum (i, mk_or (mk_and t1 (mk_not t2)) + (loop (LocalAssum (i, mk_or (mk_and t1 (mk_not t2)) (mk_and (mk_not t1) t2)) :: lit)))) ] | Kimp(t1,t2) -> @@ -1795,14 +1791,14 @@ let destructure_hyps = [mkApp (Lazy.force coq_not_imp, [| t1; t2; d1; mkVar i |])]); (onClearedName i (fun i -> - (loop (local_assum (i,mk_and t1 (mk_not t2)) :: lit)))) + (loop (LocalAssum (i,mk_and t1 (mk_not t2)) :: lit)))) ] | Kapp(Not,[t]) -> let d = decidability t in Tacticals.New.tclTHENLIST [ (generalize_tac [mkApp (Lazy.force coq_not_not, [| t; d; mkVar i |])]); - (onClearedName i (fun i -> (loop (local_assum (i,t) :: lit)))) + (onClearedName i (fun i -> (loop (LocalAssum (i,t) :: lit)))) ] | Kapp(op,[t1;t2]) -> (try @@ -1835,12 +1831,12 @@ let destructure_hyps = match destructurate_type sigma (pf_nf typ) with | Kapp(Nat,_) -> (Tacticals.New.tclTHEN - (convert_hyp_no_check (NamedDecl.set_type (EConstr.Unsafe.to_constr (mkApp (Lazy.force coq_neq, [| t1;t2|]))) + (convert_hyp_no_check (NamedDecl.set_type (mkApp (Lazy.force coq_neq, [| t1;t2|])) decl)) (loop lit)) | Kapp(Z,_) -> (Tacticals.New.tclTHEN - (convert_hyp_no_check (NamedDecl.set_type (EConstr.Unsafe.to_constr (mkApp (Lazy.force coq_Zne, [| t1;t2|]))) + (convert_hyp_no_check (NamedDecl.set_type (mkApp (Lazy.force coq_Zne, [| t1;t2|])) decl)) (loop lit)) | _ -> loop lit diff --git a/plugins/rtauto/refl_tauto.ml b/plugins/rtauto/refl_tauto.ml index 475380512..adb11f4f8 100644 --- a/plugins/rtauto/refl_tauto.ml +++ b/plugins/rtauto/refl_tauto.ml @@ -66,19 +66,18 @@ let l_E_Or = lazy (constant "E_Or") let l_D_Or = lazy (constant "D_Or") -let special_whd gl= - let infos=CClosure.create_clos_infos CClosure.all (pf_env gl) in - (fun t -> CClosure.whd_val infos (CClosure.inject t)) +let special_whd gl c = + Reductionops.clos_whd_flags CClosure.all (pf_env gl) (Tacmach.project gl) c -let special_nf gl= - let infos=CClosure.create_clos_infos CClosure.betaiotazeta (pf_env gl) in - (fun t -> CClosure.norm_val infos (CClosure.inject t)) +let special_nf gl c = + Reductionops.clos_norm_flags CClosure.betaiotazeta (pf_env gl) (Tacmach.project gl) c type atom_env= {mutable next:int; mutable env:(constr*int) list} let make_atom atom_env term= + let term = EConstr.Unsafe.to_constr term in try let (_,i)= List.find (fun (t,_)-> eq_constr term t) atom_env.env @@ -90,13 +89,17 @@ let make_atom atom_env term= Atom i let rec make_form atom_env gls term = + let open EConstr in + let open Vars in let normalize=special_nf gls in let cciterm=special_whd gls term in - match kind_of_term cciterm with + let sigma = Tacmach.project gls in + let inj = EConstr.Unsafe.to_constr in + match EConstr.kind sigma cciterm with Prod(_,a,b) -> - if EConstr.Vars.noccurn (Tacmach.project gls) 1 (EConstr.of_constr b) && + if noccurn sigma 1 b && Retyping.get_sort_family_of - (pf_env gls) (Tacmach.project gls) (EConstr.of_constr a) == InProp + (pf_env gls) sigma a == InProp then let fa=make_form atom_env gls a in let fb=make_form atom_env gls b in @@ -113,7 +116,7 @@ let rec make_form atom_env gls term = | App(hd,argv) when Int.equal (Array.length argv) 2 -> begin try - let ind, _ = destInd hd in + let ind, _ = destInd sigma hd in if Names.eq_ind ind (fst (Lazy.force li_and)) then let fa=make_form atom_env gls argv.(0) in let fb=make_form atom_env gls argv.(1) in @@ -134,9 +137,9 @@ let rec make_hyps atom_env gls lenv = function | LocalAssum (id,typ)::rest -> let hrec= make_hyps atom_env gls (typ::lenv) rest in - if List.exists (fun c -> Termops.local_occur_var Evd.empty (** FIXME *) id (EConstr.of_constr c)) lenv || + if List.exists (fun c -> Termops.local_occur_var Evd.empty (** FIXME *) id c) lenv || (Retyping.get_sort_family_of - (pf_env gls) (Tacmach.project gls) (EConstr.of_constr typ) != InProp) + (pf_env gls) (Tacmach.project gls) typ != InProp) then hrec else @@ -264,7 +267,6 @@ let rtauto_tac gls= if Retyping.get_sort_family_of (pf_env gls) (Tacmach.project gls) gl != InProp then user_err ~hdr:"rtauto" (Pp.str "goal should be in Prop") in - let gl = EConstr.Unsafe.to_constr gl in let glf=make_form gamma gls gl in let hyps=make_hyps gamma gls [gl] (pf_hyps gls) in let formula= diff --git a/plugins/rtauto/refl_tauto.mli b/plugins/rtauto/refl_tauto.mli index 9a14ac6c7..092552364 100644 --- a/plugins/rtauto/refl_tauto.mli +++ b/plugins/rtauto/refl_tauto.mli @@ -12,13 +12,13 @@ type atom_env= mutable env:(Term.constr*int) list} val make_form : atom_env -> - Proof_type.goal Tacmach.sigma -> Term.types -> Proof_search.form + Proof_type.goal Tacmach.sigma -> EConstr.types -> Proof_search.form val make_hyps : atom_env -> Proof_type.goal Tacmach.sigma -> - Term.types list -> - Context.Named.t -> + EConstr.types list -> + EConstr.named_context -> (Names.Id.t * Proof_search.form) list val rtauto_tac : Proof_type.tactic diff --git a/plugins/ssrmatching/ssrmatching.ml4 b/plugins/ssrmatching/ssrmatching.ml4 index bf3e2ac1c..d4579c3a1 100644 --- a/plugins/ssrmatching/ssrmatching.ml4 +++ b/plugins/ssrmatching/ssrmatching.ml4 @@ -838,7 +838,7 @@ let rec uniquize = function | Context.Rel.Declaration.LocalAssum _ as x -> x | Context.Rel.Declaration.LocalDef (x,_,y) -> Context.Rel.Declaration.LocalAssum(x,y) in - Environ.push_rel ctx_item env, h' + 1 in + EConstr.push_rel ctx_item env, h' + 1 in let self acc c = EConstr.of_constr (subst_loop acc (EConstr.Unsafe.to_constr c)) in let f = EConstr.of_constr f in let f' = map_constr_with_binders_left_to_right sigma inc_h self acc f in @@ -1091,7 +1091,7 @@ let thin id sigma goal = | None -> sigma | Some (hyps, concl) -> let sigma = !evdref in - let (gl,ev,sigma) = Goal.V82.mk_goal sigma hyps (EConstr.of_constr concl) (Goal.V82.extra sigma goal) in + let (gl,ev,sigma) = Goal.V82.mk_goal sigma hyps concl (Goal.V82.extra sigma goal) in let sigma = Goal.V82.partial_solution_to sigma goal gl ev in sigma diff --git a/pretyping/cases.ml b/pretyping/cases.ml index 3b5662a24..a5a5fe6d2 100644 --- a/pretyping/cases.ml +++ b/pretyping/cases.ml @@ -15,12 +15,12 @@ open Names open Nameops open Term open Termops +open Environ open EConstr open Vars open Namegen open Declarations open Inductiveops -open Environ open Reductionops open Type_errors open Glob_term @@ -38,14 +38,6 @@ open Context.Rel.Declaration module RelDecl = Context.Rel.Declaration module NamedDecl = Context.Named.Declaration -let local_assum (na, t) = - let inj = EConstr.Unsafe.to_constr in - RelDecl.LocalAssum (na, inj t) - -let local_def (na, b, t) = - let inj = EConstr.Unsafe.to_constr in - RelDecl.LocalDef (na, inj b, inj t) - (* Pattern-matching errors *) type pattern_matching_error = @@ -150,7 +142,7 @@ type tomatch_status = | Pushed of (bool*((constr * tomatch_type) * int list * Name.t)) | Alias of (bool*(Name.t * constr * (constr * types))) | NonDepAlias - | Abstract of int * Context.Rel.Declaration.t + | Abstract of int * rel_declaration type tomatch_stack = tomatch_status list @@ -261,7 +253,7 @@ type 'a pattern_matching_problem = mat : 'a matrix; caseloc : Loc.t; casestyle : case_style; - typing_function: type_constraint -> env -> evar_map ref -> 'a option -> EConstr.unsafe_judgment } + typing_function: type_constraint -> env -> evar_map ref -> 'a option -> unsafe_judgment } (*--------------------------------------------------------------------------* * A few functions to infer the inductive type from the patterns instead of * @@ -331,14 +323,12 @@ let binding_vars_of_inductive sigma = function let extract_inductive_data env sigma decl = match decl with | LocalAssum (_,t) -> - let t = EConstr.of_constr t in let tmtyp = try try_find_ind env sigma t None with Not_found -> NotInd (None,t) in let tmtypvars = binding_vars_of_inductive sigma tmtyp in (tmtyp,tmtypvars) | LocalDef (_,_,t) -> - let t = EConstr.of_constr t in (NotInd (None, t), []) let unify_tomatch_with_patterns evdref env loc typ pats realnames = @@ -451,7 +441,7 @@ let remove_current_pattern eqn = let push_current_pattern (cur,ty) eqn = match eqn.patterns with | pat::pats -> - let rhs_env = push_rel (local_def (alias_of_pat pat,cur,ty)) eqn.rhs.rhs_env in + let rhs_env = push_rel (LocalDef (alias_of_pat pat,cur,ty)) eqn.rhs.rhs_env in { eqn with rhs = { eqn.rhs with rhs_env = rhs_env }; patterns = pats } @@ -553,8 +543,8 @@ let dependencies_in_pure_rhs nargs eqns = let dependent_decl sigma a = function - | LocalAssum (na,t) -> dependent sigma a (EConstr.of_constr t) - | LocalDef (na,c,t) -> dependent sigma a (EConstr.of_constr t) || dependent sigma a (EConstr.of_constr c) + | LocalAssum (na,t) -> dependent sigma a t + | LocalDef (na,c,t) -> dependent sigma a t || dependent sigma a c let rec dep_in_tomatch sigma n = function | (Pushed _ | Alias _ | NonDepAlias) :: l -> dep_in_tomatch sigma n l @@ -625,7 +615,7 @@ let relocate_index_tomatch sigma n1 n2 = NonDepAlias :: genrec depth rest | Abstract (i,d) :: rest -> let i = relocate_rel n1 n2 depth i in - Abstract (i, RelDecl.map_constr (fun c -> EConstr.Unsafe.to_constr (relocate_index sigma n1 n2 depth (EConstr.of_constr c))) d) + Abstract (i, RelDecl.map_constr (fun c -> relocate_index sigma n1 n2 depth c) d) :: genrec (depth+1) rest in genrec 0 @@ -658,7 +648,7 @@ let replace_tomatch sigma n c = | NonDepAlias :: rest -> NonDepAlias :: replrec depth rest | Abstract (i,d) :: rest -> - Abstract (i, RelDecl.map_constr (fun t -> EConstr.Unsafe.to_constr (replace_term sigma n c depth (EConstr.of_constr t))) d) + Abstract (i, RelDecl.map_constr (fun t -> replace_term sigma n c depth t) d) :: replrec (depth+1) rest in replrec 0 @@ -683,7 +673,7 @@ let rec liftn_tomatch_stack n depth = function NonDepAlias :: liftn_tomatch_stack n depth rest | Abstract (i,d)::rest -> let i = if i let na = merge_name - (fun (LocalAssum (na,t) | LocalDef (na,_,t)) -> Name (next_name_away (named_hd env t na) avoid)) + (fun (LocalAssum (na,t) | LocalDef (na,_,t)) -> Name (next_name_away (named_hd env (EConstr.Unsafe.to_constr t) na) avoid)) d na in (na::l,(out_name na)::avoid)) @@ -1145,7 +1135,7 @@ let postprocess_dependencies evd tocheck brs tomatch pred deps cs = let rec aux k brs tomatch pred tocheck deps = match deps, tomatch with | [], _ -> brs,tomatch,pred,[] | n::deps, Abstract (i,d) :: tomatch -> - let d = map_constr (fun c -> EConstr.Unsafe.to_constr (nf_evar evd (EConstr.of_constr c))) d in + let d = map_constr (fun c -> nf_evar evd c) d in let is_d = match d with LocalAssum _ -> false | LocalDef _ -> true in if is_d || List.exists (fun c -> dependent_decl evd (lift k c) d) tocheck && Array.exists (is_dependent_branch evd k) brs then @@ -1215,12 +1205,12 @@ let rec generalize_problem names pb = function | [] -> pb, [] | i::l -> let pb',deps = generalize_problem names pb l in - let d = map_constr (CVars.lift i) (Environ.lookup_rel i pb.env) in + let d = map_constr (lift i) (lookup_rel i pb.env) in begin match d with | LocalDef (Anonymous,_,_) -> pb', deps | _ -> (* for better rendering *) - let d = RelDecl.map_type (fun c -> EConstr.Unsafe.to_constr (whd_betaiota !(pb.evdref) (EConstr.of_constr c))) d in + let d = RelDecl.map_type (fun c -> whd_betaiota !(pb.evdref) c) d in let tomatch = lift_tomatch_stack 1 pb'.tomatch in let tomatch = relocate_index_tomatch !(pb.evdref) (i+1) 1 tomatch in { pb' with @@ -1247,6 +1237,7 @@ let build_branch initial current realargs deps (realnames,curname) pb arsign eqn (* build the name x1..xn from the names present in the equations *) (* that had matched constructor C *) let cs_args = const_info.cs_args in + let cs_args = List.map (fun d -> map_rel_decl EConstr.of_constr d) cs_args in let names,aliasname = get_names pb.env cs_args eqns in let typs = List.map2 RelDecl.set_name names cs_args in @@ -1259,7 +1250,7 @@ let build_branch initial current realargs deps (realnames,curname) pb arsign eqn (* We adjust the terms to match in the context they will be once the *) (* context [x1:T1,..,xn:Tn] will have been pushed on the current env *) let typs' = - List.map_i (fun i d -> (mkRel i, map_constr (CVars.lift i) d)) 1 typs in + List.map_i (fun i d -> (mkRel i, map_constr (lift i) d)) 1 typs in let extenv = push_rel_context typs pb.env in @@ -1415,7 +1406,7 @@ and shift_problem ((current,t),_,na) pb = let pred = specialize_predicate_var (current,t,na) pb.tomatch pb.pred in let pb = { pb with - env = push_rel (local_def (na,current,ty)) pb.env; + env = push_rel (LocalDef (na,current,ty)) pb.env; tomatch = tomatch; pred = lift_predicate 1 pred tomatch; history = pop_history pb.history; @@ -1463,7 +1454,7 @@ and compile_generalization pb i d rest = ([false]). *) and compile_alias initial pb (na,orig,(expanded,expanded_typ)) rest = let f c t = - let alias = local_def (na,c,t) in + let alias = LocalDef (na,c,t) in let pb = { pb with env = push_rel alias pb.env; @@ -1601,7 +1592,7 @@ let adjust_to_extended_env_and_remove_deps env extenv sigma subst t = let (p, _, _) = lookup_rel_id x (rel_context extenv) in let rec traverse_local_defs p = match lookup_rel p extenv with - | LocalDef (_,c,_) -> assert (isRel sigma (EConstr.of_constr c)); traverse_local_defs (p + destRel sigma (EConstr.of_constr c)) + | LocalDef (_,c,_) -> assert (isRel sigma c); traverse_local_defs (p + destRel sigma c) | LocalAssum _ -> p in let p = traverse_local_defs p in let u = lift (n' - n) u in @@ -1743,6 +1734,7 @@ let build_inversion_problem loc env sigma tms t = let pat,acc = make_patvar t acc in let indf' = lift_inductive_family n indf in let sign = make_arity_signature env true indf' in + let sign = List.map (fun d -> map_rel_decl EConstr.of_constr d) sign in let patl = pat :: List.rev patl in let patl,sign = recover_and_adjust_alias_names patl sign in let p = List.length patl in @@ -1751,7 +1743,7 @@ let build_inversion_problem loc env sigma tms t = List.rev_append patl patl',acc_sign,acc | (t, NotInd (bo,typ)) :: tms -> let pat,acc = make_patvar t acc in - let d = local_assum (alias_of_pat pat,typ) in + let d = LocalAssum (alias_of_pat pat,typ) in let patl,acc_sign,acc = aux (n+1) (push_rel d env) (d::acc_sign) tms acc in pat::patl,acc_sign,acc in let avoid0 = ids_of_context env in @@ -1768,7 +1760,7 @@ let build_inversion_problem loc env sigma tms t = let n = List.length sign in let decls = - List.map_i (fun i d -> (mkRel i, map_constr (CVars.lift i) d)) 1 sign in + List.map_i (fun i d -> (mkRel i, map_constr (lift i) d)) 1 sign in let pb_env = push_rel_context sign env in let decls = @@ -1855,8 +1847,8 @@ let extract_arity_signature ?(dolift=true) env0 tomatchl tmsign = | NotInd (bo,typ) -> (match t with | None -> (match bo with - | None -> [local_assum (na, lift n typ)] - | Some b -> [local_def (na, lift n b, lift n typ)]) + | None -> [LocalAssum (na, lift n typ)] + | Some b -> [LocalDef (na, lift n b, lift n typ)]) | Some (loc,_,_) -> user_err ~loc (str"Unexpected type annotation for a term of non inductive type.")) @@ -1865,6 +1857,7 @@ let extract_arity_signature ?(dolift=true) env0 tomatchl tmsign = let ((ind,u),_) = dest_ind_family indf' in let nrealargs_ctxt = inductive_nrealdecls_env env0 ind in let arsign = fst (get_arity env0 indf') in + let arsign = List.map (fun d -> map_rel_decl EConstr.of_constr d) arsign in let realnal = match t with | Some (loc,ind',realnal) -> @@ -1874,7 +1867,7 @@ let extract_arity_signature ?(dolift=true) env0 tomatchl tmsign = anomaly (Pp.str "Ill-formed 'in' clause in cases"); List.rev realnal | None -> List.make nrealargs_ctxt Anonymous in - LocalAssum (na, build_dependent_inductive env0 indf') + LocalAssum (na, EConstr.of_constr (build_dependent_inductive env0 indf')) ::(List.map2 RelDecl.set_name realnal arsign) in let rec buildrec n = function | [],[] -> [] @@ -2069,7 +2062,7 @@ let constr_of_pat env evdref arsign pat avoid = let previd, id = prime avoid (Name (Id.of_string "wildcard")) in Name id, id :: avoid in - (PatVar (l, name), [local_assum (name, ty)] @ realargs, mkRel 1, ty, + (PatVar (l, name), [LocalAssum (name, ty)] @ realargs, mkRel 1, ty, (List.map (fun x -> mkRel 1) realargs), 1, avoid) | PatCstr (l,((_, i) as cstr),args,alias) -> let cind = inductive_of_constructor cstr in @@ -2110,7 +2103,7 @@ let constr_of_pat env evdref arsign pat avoid = Anonymous -> pat', sign, app, apptype, realargs, n, avoid | Name id -> - let sign = local_assum (alias, lift m ty) :: sign in + let sign = LocalAssum (alias, lift m ty) :: sign in let avoid = id :: avoid in let sign, i, avoid = try @@ -2122,14 +2115,14 @@ let constr_of_pat env evdref arsign pat avoid = (lift 1 app) (* aliased term *) in let neq = eq_id avoid id in - local_def (Name neq, mkRel 0, eq_t) :: sign, 2, neq :: avoid + LocalDef (Name neq, mkRel 0, eq_t) :: sign, 2, neq :: avoid with Reduction.NotConvertible -> sign, 1, avoid in (* Mark the equality as a hole *) pat', sign, lift i app, lift i apptype, realargs, n + i, avoid in - let pat', sign, patc, patty, args, z, avoid = typ env (EConstr.of_constr (RelDecl.get_type (List.hd arsign)), List.tl arsign) pat avoid in - pat', (sign, patc, (EConstr.of_constr (RelDecl.get_type (List.hd arsign)), args), pat'), avoid + let pat', sign, patc, patty, args, z, avoid = typ env (RelDecl.get_type (List.hd arsign), List.tl arsign) pat avoid in + pat', (sign, patc, (RelDecl.get_type (List.hd arsign), args), pat'), avoid (* shadows functional version *) @@ -2147,14 +2140,14 @@ match EConstr.kind sigma t with let rels_of_patsign sigma = List.map (fun decl -> match decl with - | LocalDef (na,t',t) when is_topvar sigma (EConstr.of_constr t') -> LocalAssum (na,t) + | LocalDef (na,t',t) when is_topvar sigma t' -> LocalAssum (na,t) | _ -> decl) let vars_of_ctx sigma ctx = let _, y = List.fold_right (fun decl (prev, vars) -> match decl with - | LocalDef (na,t',t) when is_topvar sigma (EConstr.of_constr t') -> + | LocalDef (na,t',t) when is_topvar sigma t' -> prev, (GApp (Loc.ghost, (GRef (Loc.ghost, delayed_force coq_eq_refl_ref, None)), @@ -2174,6 +2167,9 @@ let rec is_included x y = if Int.equal i i' then List.for_all2 is_included args args' else false +let lift_rel_context n l = + map_rel_context_with_binders (liftn n) l + (* liftsign is the current pattern's complete signature length. Hence pats is already typed in its full signature. However prevpatterns are in the original one signature per pattern form. @@ -2269,7 +2265,7 @@ let constrs_of_pats typing_fun env evdref eqns tomatchs sign neqs arity = match ineqs with | None -> [], arity | Some ineqs -> - [local_assum (Anonymous, ineqs)], lift 1 arity + [LocalAssum (Anonymous, ineqs)], lift 1 arity in let eqs_rels, arity = decompose_prod_n_assum !evdref neqs arity in eqs_rels @ neqs_rels @ rhs_rels', arity @@ -2280,7 +2276,7 @@ let constrs_of_pats typing_fun env evdref eqns tomatchs sign neqs arity = and btype = it_mkProd_or_LetIn j.uj_type rhs_rels' in let _btype = evd_comb1 (Typing.type_of env) evdref bbody in let branch_name = Id.of_string ("program_branch_" ^ (string_of_int !i)) in - let branch_decl = local_def (Name branch_name, lift !i bbody, lift !i btype) in + let branch_decl = LocalDef (Name branch_name, lift !i bbody, lift !i btype) in let branch = let bref = GVar (Loc.ghost, branch_name) in match vars_of_ctx !evdref rhs_rels with @@ -2329,7 +2325,7 @@ let abstract_tomatch env sigma tomatchs tycon = (fun t -> subst_term sigma (lift 1 c) (lift 1 t)) tycon in let name = next_ident_away (Id.of_string "filtered_var") names in (mkRel 1, lift_tomatch_type (succ lenctx) t) :: lift_ctx 1 prev, - local_def (Name name, lift lenctx c, lift lenctx $ type_of_tomatch t) :: ctx, + LocalDef (Name name, lift lenctx c, lift lenctx $ type_of_tomatch t) :: ctx, name :: names, tycon) ([], [], [], tycon) tomatchs in List.rev prev, ctx, tycon @@ -2356,14 +2352,12 @@ let build_dependent_signature env evdref avoid tomatchs arsign = let app_decl = List.hd arsign in (* The matched argument *) let appn = RelDecl.get_name app_decl in let appt = RelDecl.get_type app_decl in - let appt = EConstr.of_constr appt in let argsign = List.rev argsign in (* arguments in application order *) let env', nargeqs, argeqs, refl_args, slift, argsign' = List.fold_left2 (fun (env, nargeqs, argeqs, refl_args, slift, argsign') arg decl -> let name = RelDecl.get_name decl in let t = RelDecl.get_type decl in - let t = EConstr.of_constr t in let argt = Retyping.get_type_of env !evdref arg in let eq, refl_arg = if Reductionops.is_conv env !evdref argt t then @@ -2387,7 +2381,7 @@ let build_dependent_signature env evdref avoid tomatchs arsign = make_prime avoid name in (env, succ nargeqs, - (local_assum (Name (eq_id avoid previd), eq)) :: argeqs, + (LocalAssum (Name (eq_id avoid previd), eq)) :: argeqs, refl_arg :: refl_args, pred slift, RelDecl.set_name (Name id) decl :: argsign')) @@ -2401,7 +2395,7 @@ let build_dependent_signature env evdref avoid tomatchs arsign = in let refl_eq = mk_JMeq_refl evdref ty tm in let previd, id = make_prime avoid appn in - ((local_assum (Name (eq_id avoid previd), eq) :: argeqs) :: eqs, + ((LocalAssum (Name (eq_id avoid previd), eq) :: argeqs) :: eqs, succ nargeqs, refl_eq :: refl_args, pred slift, @@ -2417,7 +2411,7 @@ let build_dependent_signature env evdref avoid tomatchs arsign = mk_eq evdref (lift nar tomatch_ty) (mkRel slift) (lift nar tm) in - ([local_assum (Name (eq_id avoid previd), eq)] :: eqs, succ neqs, + ([LocalAssum (Name (eq_id avoid previd), eq)] :: eqs, succ neqs, (mk_eq_refl evdref tomatch_ty tm) :: refl_args, pred slift, (arsign' :: []) :: arsigns)) ([], 0, [], nar, []) tomatchs arsign @@ -2491,9 +2485,9 @@ let compile_program_cases loc style (typing_function, evdref) tycon env (* We push the initial terms to match and push their alias to rhs' envs *) (* names of aliases will be recovered from patterns (hence Anonymous here) *) - let out_tmt na = function NotInd (None,t) -> local_assum (na,t) - | NotInd (Some b, t) -> local_def (na,b,t) - | IsInd (typ,_,_) -> local_assum (na,typ) in + let out_tmt na = function NotInd (None,t) -> LocalAssum (na,t) + | NotInd (Some b, t) -> LocalDef (na,b,t) + | IsInd (typ,_,_) -> LocalAssum (na,typ) in let typs = List.map2 (fun na (tm,tmt) -> (tm,out_tmt na tmt)) nal tomatchs in let typs = @@ -2566,9 +2560,9 @@ let compile_cases loc style (typing_fun, evdref) tycon env (predopt, tomatchl, e (* names of aliases will be recovered from patterns (hence Anonymous *) (* here) *) - let out_tmt na = function NotInd (None,t) -> local_assum (na,t) - | NotInd (Some b,t) -> local_def (na,b,t) - | IsInd (typ,_,_) -> local_assum (na,typ) in + let out_tmt na = function NotInd (None,t) -> LocalAssum (na,t) + | NotInd (Some b,t) -> LocalDef (na,b,t) + | IsInd (typ,_,_) -> LocalAssum (na,typ) in let typs = List.map2 (fun na (tm,tmt) -> (tm,out_tmt na tmt)) nal tomatchs in let typs = diff --git a/pretyping/cases.mli b/pretyping/cases.mli index 9f26ae9ce..3df2d6873 100644 --- a/pretyping/cases.mli +++ b/pretyping/cases.mli @@ -47,11 +47,11 @@ val compile_cases : val constr_of_pat : Environ.env -> Evd.evar_map ref -> - Context.Rel.Declaration.t list -> + rel_context -> Glob_term.cases_pattern -> Names.Id.t list -> Glob_term.cases_pattern * - (Context.Rel.Declaration.t list * constr * + (rel_context * constr * (types * constr list) * Glob_term.cases_pattern) * Names.Id.t list @@ -85,7 +85,7 @@ type tomatch_status = | Pushed of (bool*((constr * tomatch_type) * int list * Name.t)) | Alias of (bool * (Name.t * constr * (constr * types))) | NonDepAlias - | Abstract of int * Context.Rel.Declaration.t + | Abstract of int * rel_declaration type tomatch_stack = tomatch_status list @@ -119,6 +119,6 @@ val prepare_predicate : Loc.t -> Environ.env -> Evd.evar_map -> (types * tomatch_type) list -> - Context.Rel.t list -> + rel_context list -> constr option -> 'a option -> (Evd.evar_map * Names.name list * constr) list diff --git a/pretyping/coercion.ml b/pretyping/coercion.ml index 91f53a886..8794f238b 100644 --- a/pretyping/coercion.ml +++ b/pretyping/coercion.ml @@ -18,10 +18,10 @@ open CErrors open Util open Names open Term +open Environ open EConstr open Vars open Reductionops -open Environ open Typeops open Pretype_errors open Classops @@ -127,10 +127,6 @@ let lift_args n sign = in liftrec (List.length sign) sign -let local_assum (na, t) = - let open Context.Rel.Declaration in - LocalAssum (na, EConstr.Unsafe.to_constr t) - let mu env evdref t = let rec aux v = let v' = hnf env !evdref v in @@ -159,7 +155,7 @@ and coerce loc env evdref (x : EConstr.constr) (y : EConstr.constr) let subco () = subset_coerce env evdref x y in let dest_prod c = match Reductionops.splay_prod_n env (!evdref) 1 c with - | [LocalAssum (na,t) | LocalDef (na,_,t)], c -> (na, EConstr.of_constr t), c + | [LocalAssum (na,t) | LocalDef (na,_,t)], c -> (na, t), c | _ -> raise NoSubtacCoercion in let coerce_application typ typ' c c' l l' = @@ -212,7 +208,7 @@ and coerce loc env evdref (x : EConstr.constr) (y : EConstr.constr) let name' = Name (Namegen.next_ident_away Namegen.default_dependent_ident (Termops.ids_of_context env)) in - let env' = push_rel (local_assum (name', a')) env in + let env' = push_rel (LocalAssum (name', a')) env in let c1 = coerce_unify env' (lift 1 a') (lift 1 a) in (* env, x : a' |- c1 : lift 1 a' > lift 1 a *) let coec1 = app_opt env' evdref c1 (mkRel 1) in @@ -260,7 +256,7 @@ and coerce loc env evdref (x : EConstr.constr) (y : EConstr.constr) | _ -> raise NoSubtacCoercion in let (pb, b), (pb', b') = remove_head a pb, remove_head a' pb' in - let env' = push_rel (local_assum (Name Namegen.default_dependent_ident, a)) env in + let env' = push_rel (LocalAssum (Name Namegen.default_dependent_ident, a)) env in let c2 = coerce_unify env' b b' in match c1, c2 with | None, None -> None @@ -489,7 +485,7 @@ let rec inh_conv_coerce_to_fail loc env evd rigidonly v t c1 = | Anonymous -> Name Namegen.default_dependent_ident | _ -> name in let open Context.Rel.Declaration in - let env1 = push_rel (local_assum (name,u1)) env in + let env1 = push_rel (LocalAssum (name,u1)) env in let (evd', v1) = inh_conv_coerce_to_fail loc env1 evd rigidonly (Some (mkRel 1)) (lift 1 u1) (lift 1 t1) in diff --git a/pretyping/constr_matching.ml b/pretyping/constr_matching.ml index 55612aa66..cad21543b 100644 --- a/pretyping/constr_matching.ml +++ b/pretyping/constr_matching.ml @@ -136,14 +136,6 @@ let make_renaming ids = function end | _ -> dummy_constr -let local_assum (na, t) = - let inj = EConstr.Unsafe.to_constr in - LocalAssum (na, inj t) - -let local_def (na, b, t) = - let inj = EConstr.Unsafe.to_constr in - LocalDef (na, inj b, inj t) - let to_fix (idx, (nas, cs, ts)) = let inj = EConstr.of_constr in (idx, (nas, Array.map inj cs, Array.map inj ts)) @@ -273,15 +265,15 @@ let matches_core env sigma convert allow_partial_app allow_bound_rels sorec ctx env subst c1 c2 | PProd (na1,c1,d1), Prod(na2,c2,d2) -> - sorec ((na1,na2,c2)::ctx) (Environ.push_rel (local_assum (na2,c2)) env) + sorec ((na1,na2,c2)::ctx) (EConstr.push_rel (LocalAssum (na2,c2)) env) (add_binders na1 na2 binding_vars (sorec ctx env subst c1 c2)) d1 d2 | PLambda (na1,c1,d1), Lambda(na2,c2,d2) -> - sorec ((na1,na2,c2)::ctx) (Environ.push_rel (local_assum (na2,c2)) env) + sorec ((na1,na2,c2)::ctx) (EConstr.push_rel (LocalAssum (na2,c2)) env) (add_binders na1 na2 binding_vars (sorec ctx env subst c1 c2)) d1 d2 | PLetIn (na1,c1,d1), LetIn(na2,c2,t2,d2) -> - sorec ((na1,na2,t2)::ctx) (Environ.push_rel (local_def (na2,c2,t2)) env) + sorec ((na1,na2,t2)::ctx) (EConstr.push_rel (LocalDef (na2,c2,t2)) env) (add_binders na1 na2 binding_vars (sorec ctx env subst c1 c2)) d1 d2 | PIf (a1,b1,b1'), Case (ci,_,a2,[|b2;b2'|]) -> @@ -290,12 +282,12 @@ let matches_core env sigma convert allow_partial_app allow_bound_rels let n = Context.Rel.length ctx_b2 in let n' = Context.Rel.length ctx_b2' in if Vars.noccur_between sigma 1 n b2 && Vars.noccur_between sigma 1 n' b2' then - let f l (LocalAssum (na,t) | LocalDef (na,_,t)) = (Anonymous,na,EConstr.of_constr t)::l in + let f l (LocalAssum (na,t) | LocalDef (na,_,t)) = (Anonymous,na,t)::l in let ctx_br = List.fold_left f ctx ctx_b2 in let ctx_br' = List.fold_left f ctx ctx_b2' in let b1 = lift_pattern n b1 and b1' = lift_pattern n' b1' in - sorec ctx_br' (Environ.push_rel_context ctx_b2' env) - (sorec ctx_br (Environ.push_rel_context ctx_b2 env) + sorec ctx_br' (push_rel_context ctx_b2' env) + (sorec ctx_br (push_rel_context ctx_b2 env) (sorec ctx env subst a1 a2) b1 b2) b1' b2' else raise PatternMatchingFailure @@ -388,21 +380,21 @@ let sub_match ?(partial_app=false) ?(closed=true) env sigma pat c = | [c1; c2] -> mk_ctx (mkLambda (x, c1, c2)) | _ -> assert false in - let env' = Environ.push_rel (local_assum (x,c1)) env in + let env' = EConstr.push_rel (LocalAssum (x,c1)) env in try_aux [(env, c1); (env', c2)] next_mk_ctx next | Prod (x,c1,c2) -> let next_mk_ctx = function | [c1; c2] -> mk_ctx (mkProd (x, c1, c2)) | _ -> assert false in - let env' = Environ.push_rel (local_assum (x,c1)) env in + let env' = EConstr.push_rel (LocalAssum (x,c1)) env in try_aux [(env, c1); (env', c2)] next_mk_ctx next | LetIn (x,c1,t,c2) -> let next_mk_ctx = function | [c1; c2] -> mk_ctx (mkLetIn (x, c1, t, c2)) | _ -> assert false in - let env' = Environ.push_rel (local_def (x,c1,t)) env in + let env' = EConstr.push_rel (LocalDef (x,c1,t)) env in try_aux [(env, c1); (env', c2)] next_mk_ctx next | App (c1,lc) -> let topdown = true in diff --git a/pretyping/evarconv.ml b/pretyping/evarconv.ml index 3ae2e35e6..f5cab070e 100644 --- a/pretyping/evarconv.ml +++ b/pretyping/evarconv.ml @@ -11,12 +11,12 @@ open Util open Names open Term open Termops +open Environ open EConstr open Vars open CClosure open Reduction open Reductionops -open Environ open Recordops open Evarutil open Evardefine @@ -58,12 +58,12 @@ let eval_flexible_term ts env evd c = | Rel n -> (try match lookup_rel n env with | RelDecl.LocalAssum _ -> None - | RelDecl.LocalDef (_,v,_) -> Some (lift n (EConstr.of_constr v)) + | RelDecl.LocalDef (_,v,_) -> Some (lift n v) with Not_found -> None) | Var id -> (try if is_transparent_variable ts id then - Option.map EConstr.of_constr (env |> lookup_named id |> NamedDecl.get_value) + env |> lookup_named id |> NamedDecl.get_value else None with Not_found -> None) | LetIn (_,b,_,c) -> Some (subst1 b c) @@ -404,7 +404,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty let eta env evd onleft sk term sk' term' = assert (match sk with [] -> true | _ -> false); let (na,c1,c'1) = destLambda evd term in - let c = EConstr.to_constr evd c1 in + let c = nf_evar evd c1 in let env' = push_rel (RelDecl.LocalAssum (na,c)) env in let out1 = whd_betaiota_deltazeta_for_iota_state (fst ts) env' evd Cst_stack.empty (c'1, Stack.empty) in @@ -612,8 +612,8 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty (fun i -> evar_conv_x ts env i CUMUL t2 t1)]); (fun i -> evar_conv_x ts env i CONV b1 b2); (fun i -> - let b = EConstr.to_constr i b1 in - let t = EConstr.to_constr i t1 in + let b = nf_evar i b1 in + let t = nf_evar i t1 in let na = Nameops.name_max na1 na2 in evar_conv_x ts (push_rel (RelDecl.LocalDef (na,b,t)) env) i pbty c'1 c'2); (fun i -> exact_ise_stack2 env i (evar_conv_x ts) sk1 sk2)] @@ -730,7 +730,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty ise_and evd [(fun i -> evar_conv_x ts env i CONV c1 c2); (fun i -> - let c = EConstr.to_constr i c1 in + let c = nf_evar i c1 in let na = Nameops.name_max na1 na2 in evar_conv_x ts (push_rel (RelDecl.LocalAssum (na,c)) env) i CONV c'1 c'2)] @@ -789,7 +789,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty ise_and evd [(fun i -> evar_conv_x ts env i CONV c1 c2); (fun i -> - let c = EConstr.to_constr i c1 in + let c = nf_evar i c1 in let na = Nameops.name_max n1 n2 in evar_conv_x ts (push_rel (RelDecl.LocalAssum (na,c)) env) i pbty c'1 c'2)] diff --git a/pretyping/evardefine.ml b/pretyping/evardefine.ml index 5831d3191..faf34baf7 100644 --- a/pretyping/evardefine.ml +++ b/pretyping/evardefine.ml @@ -11,10 +11,10 @@ open Pp open Names open Term open Termops +open Environ open EConstr open Vars open Namegen -open Environ open Evd open Evarutil open Pretype_errors @@ -22,25 +22,20 @@ open Sigma.Notations module RelDecl = Context.Rel.Declaration -let nlocal_assum (na, t) = - let open Context.Named.Declaration in - let inj = EConstr.Unsafe.to_constr in - LocalAssum (na, inj t) - let new_evar_unsafe env evd ?src ?filter ?candidates ?store ?naming ?principal typ = let evd = Sigma.Unsafe.of_evar_map evd in let Sigma (evk, evd, _) = new_evar env evd ?src ?filter ?candidates ?store ?naming ?principal typ in (Sigma.to_evar_map evd, evk) let env_nf_evar sigma env = - let nf_evar c = EConstr.Unsafe.to_constr (nf_evar sigma (EConstr.of_constr c)) in + let nf_evar c = nf_evar sigma c in process_rel_context (fun d e -> push_rel (RelDecl.map_constr nf_evar d) e) env let env_nf_betaiotaevar sigma env = process_rel_context (fun d e -> - push_rel (RelDecl.map_constr (fun c -> EConstr.Unsafe.to_constr (Reductionops.nf_betaiota sigma (EConstr.of_constr c))) d) e) env + push_rel (RelDecl.map_constr (fun c -> Reductionops.nf_betaiota sigma c) d) e) env (****************************************) (* Operations on value/type constraints *) @@ -93,7 +88,7 @@ let define_pure_evar_as_product evd evk = (Sigma.to_evar_map evd1, e) in let evd2,rng = - let newenv = push_named (nlocal_assum (id, dom)) evenv in + let newenv = push_named (LocalAssum (id, dom)) evenv in let src = evar_source evk evd1 in let filter = Filter.extend 1 (evar_filter evi) in if is_prop_sort s then @@ -146,7 +141,7 @@ let define_pure_evar_as_lambda env evd evk = let avoid = ids_of_named_context (evar_context evi) in let id = next_name_away_with_default_using_types "x" na avoid (EConstr.Unsafe.to_constr (Reductionops.whd_evar evd dom)) in - let newenv = push_named (nlocal_assum (id, dom)) evenv in + let newenv = push_named (LocalAssum (id, dom)) evenv in let filter = Filter.extend 1 (evar_filter evi) in let src = evar_source evk evd1 in let evd2,body = new_evar_unsafe newenv evd1 ~src (subst1 (mkVar id) rng) ~filter in diff --git a/pretyping/evarsolve.ml b/pretyping/evarsolve.ml index 3235c2505..ff4736528 100644 --- a/pretyping/evarsolve.ml +++ b/pretyping/evarsolve.ml @@ -23,14 +23,6 @@ open Evarutil open Pretype_errors open Sigma.Notations -let nlocal_assum (na, t) = - let inj = EConstr.Unsafe.to_constr in - Context.Named.Declaration.LocalAssum (na, inj t) - -let nlocal_def (na, b, t) = - let inj = EConstr.Unsafe.to_constr in - Context.Named.Declaration.LocalDef (na, inj b, inj t) - let normalize_evar evd ev = match EConstr.kind evd (mkEvar ev) with | Evar (evk,args) -> (evk,args) @@ -264,7 +256,6 @@ let compute_var_aliases sign sigma = let id = get_id decl in match decl with | LocalDef (_,t,_) -> - let t = EConstr.of_constr t in (match EConstr.kind sigma t with | Var id' -> let aliases_of_id = @@ -281,8 +272,6 @@ let compute_rel_aliases var_aliases rels sigma = (n-1, match decl with | LocalDef (_,t,u) -> - let t = EConstr.of_constr t in - let u = EConstr.of_constr u in (match EConstr.kind sigma t with | Var id' -> let aliases_of_n = @@ -338,7 +327,6 @@ let extend_alias sigma decl (var_aliases,rel_aliases) = let rel_aliases = match decl with | LocalDef(_,t,_) -> - let t = EConstr.of_constr t in (match EConstr.kind sigma t with | Var id' -> let aliases_of_binder = @@ -530,7 +518,7 @@ let solve_pattern_eqn env sigma l c = (* Rem: if [a] links to a let-in, do as if it were an assumption *) | Rel n -> let open Context.Rel.Declaration in - let d = map_constr (CVars.lift n) (lookup_rel n env) in + let d = map_constr (lift n) (lookup_rel n env) in mkLambda_or_LetIn d c' | Var id -> let d = lookup_named id env in mkNamedLambda_or_LetIn d c' @@ -556,6 +544,7 @@ let solve_pattern_eqn env sigma l c = let make_projectable_subst aliases sigma evi args = let sign = evar_filtered_context evi in + let sign = List.map (fun d -> map_named_decl EConstr.of_constr d) sign in let evar_aliases = compute_var_aliases sign sigma in let (_,full_subst,cstr_subst) = List.fold_right @@ -571,7 +560,7 @@ let make_projectable_subst aliases sigma evi args = | _ -> cstrs in (rest,Id.Map.add id [a,normalize_alias_opt sigma aliases a,id] all,cstrs) | LocalDef (id,c,_), a::rest -> - (match EConstr.kind sigma (EConstr.of_constr c) with + (match EConstr.kind sigma c with | Var id' -> let idc = normalize_alias_var sigma evar_aliases id' in let sub = try Id.Map.find idc all with Not_found -> [] in @@ -646,19 +635,17 @@ let materialize_evar define_fun env evd k (evk1,args1) ty_in_env = let LocalAssum (na,t_in_env) | LocalDef (na,_,t_in_env) = d in let id = next_name_away na avoid in let evd,t_in_sign = - let t_in_env = EConstr.of_constr t_in_env in let s = Retyping.get_sort_of env evd t_in_env in let evd,ty_t_in_sign = refresh_universes ~status:univ_flexible (Some false) env evd (mkSort s) in define_evar_from_virtual_equation define_fun env evd src t_in_env ty_t_in_sign sign filter inst_in_env in let evd,d' = match d with - | LocalAssum _ -> evd, nlocal_assum (id,t_in_sign) + | LocalAssum _ -> evd, Context.Named.Declaration.LocalAssum (id,t_in_sign) | LocalDef (_,b,_) -> - let b = EConstr.of_constr b in let evd,b = define_evar_from_virtual_equation define_fun env evd src b t_in_sign sign filter inst_in_env in - evd, nlocal_def (id,b,t_in_sign) in + evd, Context.Named.Declaration.LocalDef (id,b,t_in_sign) in (push_named_context_val d' sign, Filter.extend 1 filter, (mkRel 1)::(List.map (lift 1) inst_in_env), (mkRel 1)::(List.map (lift 1) inst_in_sign), @@ -1238,9 +1225,11 @@ let solve_evar_evar ?(force=false) f g env evd pbty (evk1,args1 as ev1) (evk2,ar The body of ?X and ?Y just has to be of type Π Δ. Type k for some k <= i, j. *) let evienv = Evd.evar_env evi in let ctx1, i = Reduction.dest_arity evienv evi.evar_concl in + let ctx1 = List.map (fun c -> map_rel_decl EConstr.of_constr c) ctx1 in let evi2 = Evd.find evd evk2 in let evi2env = Evd.evar_env evi2 in let ctx2, j = Reduction.dest_arity evi2env evi2.evar_concl in + let ctx2 = List.map (fun c -> map_rel_decl EConstr.of_constr c) ctx2 in let ui, uj = univ_of_sort i, univ_of_sort j in if i == j || Evd.check_eq evd ui uj then (* Shortcut, i = j *) @@ -1397,7 +1386,7 @@ let rec invert_definition conv_algo choose env evd pbty (evk,argsv as ev) rhs = (* No unique projection but still restrict to where it is possible *) (* materializing is necessary, but is restricting useful? *) let ty = find_solution_type (evar_filtered_env evi) sols in - let ty' = instantiate_evar_array evi (EConstr.of_constr ty) argsv in + let ty' = instantiate_evar_array evi ty argsv in let (evd,evar,(evk',argsv' as ev')) = materialize_evar (evar_define conv_algo ~choose) env !evdref 0 ev ty' in let ts = expansions_of_var evd aliases t in diff --git a/pretyping/find_subterm.ml b/pretyping/find_subterm.ml index d09686f6e..3fc569fc4 100644 --- a/pretyping/find_subterm.ml +++ b/pretyping/find_subterm.ml @@ -64,8 +64,8 @@ let proceed_with_occurrences f occs x = let map_named_declaration_with_hyploc f hyploc acc decl = let open Context.Named.Declaration in let f acc typ = - let acc, typ = f (Some (NamedDecl.get_id decl, hyploc)) acc (EConstr.of_constr typ) in - acc, EConstr.Unsafe.to_constr typ + let acc, typ = f (Some (NamedDecl.get_id decl, hyploc)) acc typ in + acc, typ in match decl,hyploc with | LocalAssum (id,_), InHypValueOnly -> diff --git a/pretyping/find_subterm.mli b/pretyping/find_subterm.mli index 3d2ebb72d..e3d3b74f1 100644 --- a/pretyping/find_subterm.mli +++ b/pretyping/find_subterm.mli @@ -51,7 +51,7 @@ val replace_term_occ_decl_modulo : evar_map -> (occurrences * hyp_location_flag) or_like_first -> 'a testing_function -> (unit -> constr) -> - Context.Named.Declaration.t -> Context.Named.Declaration.t + named_declaration -> named_declaration (** [subst_closed_term_occ occl c d] replaces occurrences of closed [c] at positions [occl] by [Rel 1] in [d] (see also Note OCC), @@ -63,7 +63,7 @@ val subst_closed_term_occ : env -> evar_map -> occurrences or_like_first -> closed [c] at positions [occl] by [Rel 1] in [decl]. *) val subst_closed_term_occ_decl : env -> evar_map -> (occurrences * hyp_location_flag) or_like_first -> - constr -> Context.Named.Declaration.t -> Context.Named.Declaration.t * evar_map + constr -> named_declaration -> named_declaration * evar_map (** Miscellaneous *) val error_invalid_occurrence : int list -> 'a diff --git a/pretyping/inductiveops.ml b/pretyping/inductiveops.ml index c00ceb02e..3191a58ff 100644 --- a/pretyping/inductiveops.ml +++ b/pretyping/inductiveops.ml @@ -371,7 +371,7 @@ let make_case_or_project env sigma indf ci pred c branches = | LocalAssum (na, t) -> let t = mkProj (Projection.make ps.(i) true, c) in (i + 1, t :: subst) - | LocalDef (na, b, t) -> (i, Vars.substl subst (EConstr.of_constr b) :: subst)) + | LocalDef (na, b, t) -> (i, Vars.substl subst b :: subst)) ctx (0, []) in Vars.substl subst br diff --git a/pretyping/patternops.ml b/pretyping/patternops.ml index 26e23be23..954aa6a94 100644 --- a/pretyping/patternops.ml +++ b/pretyping/patternops.ml @@ -121,19 +121,10 @@ let head_of_constr_reference sigma c = match EConstr.kind sigma c with | Var id -> VarRef id | _ -> anomaly (Pp.str "Not a rigid reference") -let local_assum (na, t) = - let open Context.Rel.Declaration in - let inj = EConstr.Unsafe.to_constr in - LocalAssum (na, inj t) - -let local_def (na, b, t) = - let open Context.Rel.Declaration in - let inj = EConstr.Unsafe.to_constr in - LocalDef (na, inj b, inj t) - let pattern_of_constr env sigma t = let open EConstr in let rec pattern_of_constr env t = + let open Context.Rel.Declaration in match EConstr.kind sigma t with | Rel n -> PRel n | Meta n -> PMeta (Some (Id.of_string ("META" ^ string_of_int n))) @@ -143,11 +134,11 @@ let pattern_of_constr env sigma t = | Sort (Type _) -> PSort (GType []) | Cast (c,_,_) -> pattern_of_constr env c | LetIn (na,c,t,b) -> PLetIn (na,pattern_of_constr env c, - pattern_of_constr (push_rel (local_def (na,c,t)) env) b) + pattern_of_constr (push_rel (LocalDef (na,c,t)) env) b) | Prod (na,c,b) -> PProd (na,pattern_of_constr env c, - pattern_of_constr (push_rel (local_assum (na, c)) env) b) + pattern_of_constr (push_rel (LocalAssum (na, c)) env) b) | Lambda (na,c,b) -> PLambda (na,pattern_of_constr env c, - pattern_of_constr (push_rel (local_assum (na, c)) env) b) + pattern_of_constr (push_rel (LocalAssum (na, c)) env) b) | App (f,a) -> (match match EConstr.kind sigma f with diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml index 4660978df..2470decdd 100644 --- a/pretyping/pretyping.ml +++ b/pretyping/pretyping.ml @@ -28,10 +28,10 @@ open Names open Evd open Term open Termops +open Environ open EConstr open Vars open Reductionops -open Environ open Type_errors open Typeops open Typing @@ -70,16 +70,6 @@ open Inductiveops (************************************************************************) -let local_assum (na, t) = - let open Context.Rel.Declaration in - let inj = EConstr.Unsafe.to_constr in - LocalAssum (na, inj t) - -let local_def (na, b, t) = - let open Context.Rel.Declaration in - let inj = EConstr.Unsafe.to_constr in - LocalDef (na, inj b, inj t) - module ExtraEnv = struct @@ -94,7 +84,7 @@ let get_extra env = let ids = List.map get_id (named_context env) in let avoid = List.fold_right Id.Set.add ids Id.Set.empty in Context.Rel.fold_outside push_rel_decl_to_named_context - (Environ.rel_context env) ~init:(empty_csubst, [], avoid, named_context env) + (rel_context env) ~init:(empty_csubst, [], avoid, named_context env) let make_env env = { env = env; extra = lazy (get_extra env) } let rel_context env = rel_context env.env @@ -116,7 +106,7 @@ let lookup_named id env = lookup_named id env.env let e_new_evar env evdref ?src ?naming typ = let subst2 subst vsubst c = csubst_subst subst (replace_vars vsubst c) in let open Context.Named.Declaration in - let inst_vars = List.map (get_id %> EConstr.mkVar) (named_context env.env) in + let inst_vars = List.map (get_id %> mkVar) (named_context env.env) in let inst_rels = List.rev (rel_list 0 (nb_rel env.env)) in let (subst, vsubst, _, nc) = Lazy.force env.extra in let typ' = subst2 subst vsubst typ in @@ -128,7 +118,7 @@ let e_new_evar env evdref ?src ?naming typ = e let push_rec_types (lna,typarray,_) env = - let ctxt = Array.map2_i (fun i na t -> local_assum (na, lift i t)) lna typarray in + let ctxt = Array.map2_i (fun i na t -> Context.Rel.Declaration.LocalAssum (na, lift i t)) lna typarray in Array.fold_left (fun e assum -> push_rel assum e) env ctxt end @@ -434,7 +424,6 @@ let pretype_id pretype k0 loc env evdref lvar id = (* Look for the binder of [id] *) try let (n,_,typ) = lookup_rel_id id (rel_context env) in - let typ = EConstr.of_constr typ in { uj_val = mkRel n; uj_type = lift n typ } with Not_found -> let env = ltac_interp_name_env k0 lvar env in @@ -468,7 +457,7 @@ let pretype_id pretype k0 loc env evdref lvar id = end; (* Check if [id] is a section or goal variable *) try - { uj_val = mkVar id; uj_type = EConstr.of_constr (NamedDecl.get_type (lookup_named id env)) } + { uj_val = mkVar id; uj_type = NamedDecl.get_type (lookup_named id env) } with Not_found -> (* [id] not found, standard error message *) error_var_not_found ~loc id @@ -511,7 +500,7 @@ let pretype_ref loc evdref env ref us = match ref with | VarRef id -> (* Section variable *) - (try make_judge (mkVar id) (EConstr.of_constr (NamedDecl.get_type (lookup_named id env))) + (try make_judge (mkVar id) (NamedDecl.get_type (lookup_named id env)) with Not_found -> (* This may happen if env is a goal env and section variables have been cleared - section variables should be different from goal @@ -614,14 +603,14 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre [] -> ctxt | (na,bk,None,ty)::bl -> let ty' = pretype_type empty_valcon env evdref lvar ty in - let dcl = local_assum (na, ty'.utj_val) in - let dcl' = local_assum (ltac_interp_name lvar na,ty'.utj_val) in + let dcl = LocalAssum (na, ty'.utj_val) in + let dcl' = LocalAssum (ltac_interp_name lvar na,ty'.utj_val) in type_bl (push_rel dcl env) (Context.Rel.add dcl' ctxt) bl | (na,bk,Some bd,ty)::bl -> let ty' = pretype_type empty_valcon env evdref lvar ty in let bd' = pretype (mk_tycon ty'.utj_val) env evdref lvar bd in - let dcl = local_def (na, bd'.uj_val, ty'.utj_val) in - let dcl' = local_def (ltac_interp_name lvar na, bd'.uj_val, ty'.utj_val) in + let dcl = LocalDef (na, bd'.uj_val, ty'.utj_val) in + let dcl' = LocalDef (ltac_interp_name lvar na, bd'.uj_val, ty'.utj_val) in type_bl (push_rel dcl env) (Context.Rel.add dcl' ctxt) bl in let ctxtv = Array.map (type_bl env Context.Rel.empty) bl in let larj = @@ -793,7 +782,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre (* The name specified by ltac is used also to create bindings. So the substitution must also be applied on variables before they are looked up in the rel context. *) - let var = local_assum (name, j.utj_val) in + let var = LocalAssum (name, j.utj_val) in let j' = pretype rng (push_rel var env) evdref lvar c2 in let name = ltac_interp_name lvar name in let resj = judge_of_abstraction env.ExtraEnv.env (orelse_name name name') j j' in @@ -809,7 +798,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre let j = pretype_type empty_valcon env evdref lvar c2 in { j with utj_val = lift 1 j.utj_val } | Name _ -> - let var = local_assum (name, j.utj_val) in + let var = LocalAssum (name, j.utj_val) in let env' = push_rel var env in pretype_type empty_valcon env' evdref lvar c2 in @@ -837,7 +826,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre (* The name specified by ltac is used also to create bindings. So the substitution must also be applied on variables before they are looked up in the rel context. *) - let var = local_def (name, j.uj_val, t) in + let var = LocalDef (name, j.uj_val, t) in let tycon = lift_tycon 1 tycon in let j' = pretype tycon (push_rel var env) evdref lvar c2 in let name = ltac_interp_name lvar name in @@ -861,6 +850,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre user_err ~loc:loc (str "Destructing let on this type expects " ++ int cs.cs_nargs ++ str " variables."); let fsign, record = + let set_name na d = set_name na (map_rel_decl EConstr.of_constr d) in match get_projections env.ExtraEnv.env indf with | None -> List.map2 set_name (List.rev nal) cs.cs_args, false @@ -870,7 +860,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre | na :: names, (LocalAssum (_,t) :: l) -> let t = EConstr.of_constr t in let proj = Projection.make ps.(cs.cs_nargs - k) true in - local_def (na, lift (cs.cs_nargs - n) (mkProj (proj, cj.uj_val)), t) + LocalDef (na, lift (cs.cs_nargs - n) (mkProj (proj, cj.uj_val)), t) :: aux (n+1) (k + 1) names l | na :: names, (decl :: l) -> set_name na decl :: aux (n+1) k names l @@ -896,6 +886,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre else arsgn in let psign = LocalAssum (na, build_dependent_inductive env.ExtraEnv.env indf) :: arsgn in + let psign = List.map (fun d -> map_rel_decl EConstr.of_constr d) psign in let nar = List.length arsgn in (match po with | Some p -> @@ -903,6 +894,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre let pj = pretype_type empty_valcon env_p evdref lvar p in let ccl = nf_evar !evdref pj.utj_val in let psign = make_arity_signature env.ExtraEnv.env true indf in (* with names *) + let psign = List.map (fun d -> map_rel_decl EConstr.of_constr d) psign in let p = it_mkLambda_or_LetIn ccl psign in let inst = (Array.map_to_list EConstr.of_constr cs.cs_concl_realargs) @@ -956,6 +948,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre in let nar = List.length arsgn in let psign = LocalAssum (na, build_dependent_inductive env.ExtraEnv.env indf) :: arsgn in + let psign = List.map (fun d -> map_rel_decl EConstr.of_constr d) psign in let pred,p = match po with | Some p -> let env_p = push_rel_context psign env in @@ -978,17 +971,18 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre let n = Context.Rel.length cs.cs_args in let pi = lift n pred in (* liftn n 2 pred ? *) let pi = beta_applist !evdref (pi, [EConstr.of_constr (build_dependent_constructor cs)]) in + let cs_args = List.map (fun d -> map_rel_decl EConstr.of_constr d) cs.cs_args in let csgn = if not !allow_anonymous_refs then - List.map (set_name Anonymous) cs.cs_args + List.map (set_name Anonymous) cs_args else List.map (map_name (function Name _ as n -> n | Anonymous -> Name Namegen.default_non_dependent_ident)) - cs.cs_args + cs_args in let env_c = push_rel_context csgn env in let bj = pretype (mk_tycon pi) env_c evdref lvar b in - it_mkLambda_or_LetIn bj.uj_val cs.cs_args in + it_mkLambda_or_LetIn bj.uj_val cs_args in let b1 = f cstrs.(0) b1 in let b2 = f cstrs.(1) b2 in let v = @@ -1060,12 +1054,10 @@ and pretype_instance k0 resolve_tc env evdref lvar loc hyps evk update = with Not_found -> try let (n,_,t') = lookup_rel_id id (rel_context env) in - let t' = EConstr.of_constr t' in if is_conv env.ExtraEnv.env !evdref t t' then mkRel n, update else raise Not_found with Not_found -> try let t' = env |> lookup_named id |> NamedDecl.get_type in - let t' = EConstr.of_constr t' in if is_conv env.ExtraEnv.env !evdref t t' then mkVar id, update else raise Not_found with Not_found -> user_err ~loc (str "Cannot interpret " ++ diff --git a/pretyping/reductionops.ml b/pretyping/reductionops.ml index bc5c629f4..a1585ef52 100644 --- a/pretyping/reductionops.ml +++ b/pretyping/reductionops.ml @@ -13,9 +13,9 @@ open Term open Termops open Univ open Evd +open Environ open EConstr open Vars -open Environ open Context.Rel.Declaration exception Elimconst @@ -609,14 +609,6 @@ let pr_state (tm,sk) = let pr c = Termops.print_constr c in h 0 (pr tm ++ str "|" ++ cut () ++ Stack.pr pr sk) -let local_assum (na, t) = - let inj = EConstr.Unsafe.to_constr in - LocalAssum (na, inj t) - -let local_def (na, b, t) = - let inj = EConstr.Unsafe.to_constr in - LocalDef (na, inj b, inj t) - (*************************************) (*** Reduction Functions Operators ***) (*************************************) @@ -851,12 +843,12 @@ let rec whd_state_gen ?csts ~refold ~tactic_mode flags env sigma = match c0 with | Rel n when CClosure.RedFlags.red_set flags CClosure.RedFlags.fDELTA -> (match lookup_rel n env with - | LocalDef (_,body,_) -> whrec Cst_stack.empty (lift n (EConstr.of_constr body), stack) + | LocalDef (_,body,_) -> whrec Cst_stack.empty (lift n body, stack) | _ -> fold ()) | Var id when CClosure.RedFlags.red_set flags (CClosure.RedFlags.fVAR id) -> (match lookup_named id env with | LocalDef (_,body,_) -> - whrec (if refold then Cst_stack.add_cst (mkVar id) cst_l else cst_l) (EConstr.of_constr body, stack) + whrec (if refold then Cst_stack.add_cst (mkVar id) cst_l else cst_l) (body, stack) | _ -> fold ()) | Evar ev -> fold () | Meta ev -> @@ -958,7 +950,7 @@ let rec whd_state_gen ?csts ~refold ~tactic_mode flags env sigma = | Some _ when CClosure.RedFlags.red_set flags CClosure.RedFlags.fBETA -> apply_subst (fun _ -> whrec) [] sigma refold cst_l x stack | None when CClosure.RedFlags.red_set flags CClosure.RedFlags.fETA -> - let env' = push_rel (local_assum (na, t)) env in + let env' = push_rel (LocalAssum (na, t)) env in let whrec' = whd_state_gen ~refold ~tactic_mode flags env' sigma in (match EConstr.kind sigma (Stack.zip ~refold sigma (fst (whrec' (c, Stack.empty)))) with | App (f,cl) -> @@ -1468,7 +1460,7 @@ let splay_prod env sigma = let t = whd_all env sigma c in match EConstr.kind sigma t with | Prod (n,a,c0) -> - decrec (push_rel (local_assum (n,a)) env) + decrec (push_rel (LocalAssum (n,a)) env) (bind_assum (n,a)::m) c0 | _ -> m,t in @@ -1479,7 +1471,7 @@ let splay_lam env sigma = let t = whd_all env sigma c in match EConstr.kind sigma t with | Lambda (n,a,c0) -> - decrec (push_rel (local_assum (n,a)) env) + decrec (push_rel (LocalAssum (n,a)) env) (bind_assum (n,a)::m) c0 | _ -> m,t in @@ -1490,11 +1482,11 @@ let splay_prod_assum env sigma = let t = whd_allnolet env sigma c in match EConstr.kind sigma t with | Prod (x,t,c) -> - prodec_rec (push_rel (local_assum (x,t)) env) - (Context.Rel.add (local_assum (x,t)) l) c + prodec_rec (push_rel (LocalAssum (x,t)) env) + (Context.Rel.add (LocalAssum (x,t)) l) c | LetIn (x,b,t,c) -> - prodec_rec (push_rel (local_def (x,b,t)) env) - (Context.Rel.add (local_def (x,b,t)) l) c + prodec_rec (push_rel (LocalDef (x,b,t)) env) + (Context.Rel.add (LocalDef (x,b,t)) l) c | Cast (c,_,_) -> prodec_rec env l c | _ -> let t' = whd_all env sigma t in @@ -1515,8 +1507,8 @@ let splay_prod_n env sigma n = let rec decrec env m ln c = if Int.equal m 0 then (ln,c) else match EConstr.kind sigma (whd_all env sigma c) with | Prod (n,a,c0) -> - decrec (push_rel (local_assum (n,a)) env) - (m-1) (Context.Rel.add (local_assum (n,a)) ln) c0 + decrec (push_rel (LocalAssum (n,a)) env) + (m-1) (Context.Rel.add (LocalAssum (n,a)) ln) c0 | _ -> invalid_arg "splay_prod_n" in decrec env n Context.Rel.empty @@ -1525,8 +1517,8 @@ let splay_lam_n env sigma n = let rec decrec env m ln c = if Int.equal m 0 then (ln,c) else match EConstr.kind sigma (whd_all env sigma c) with | Lambda (n,a,c0) -> - decrec (push_rel (local_assum (n,a)) env) - (m-1) (Context.Rel.add (local_assum (n,a)) ln) c0 + decrec (push_rel (LocalAssum (n,a)) env) + (m-1) (Context.Rel.add (LocalAssum (n,a)) ln) c0 | _ -> invalid_arg "splay_lam_n" in decrec env n Context.Rel.empty @@ -1566,8 +1558,8 @@ let find_conclusion env sigma = let rec decrec env c = let t = whd_all env sigma c in match EConstr.kind sigma t with - | Prod (x,t,c0) -> decrec (push_rel (local_assum (x,t)) env) c0 - | Lambda (x,t,c0) -> decrec (push_rel (local_assum (x,t)) env) c0 + | Prod (x,t,c0) -> decrec (push_rel (LocalAssum (x,t)) env) c0 + | Lambda (x,t,c0) -> decrec (push_rel (LocalAssum (x,t)) env) c0 | t -> t in decrec env diff --git a/pretyping/reductionops.mli b/pretyping/reductionops.mli index dcc11cfcf..15ddeb15c 100644 --- a/pretyping/reductionops.mli +++ b/pretyping/reductionops.mli @@ -217,10 +217,10 @@ val splay_prod : env -> evar_map -> constr -> (Name.t * constr) list * constr val splay_lam : env -> evar_map -> constr -> (Name.t * constr) list * constr val splay_arity : env -> evar_map -> constr -> (Name.t * constr) list * sorts val sort_of_arity : env -> evar_map -> constr -> sorts -val splay_prod_n : env -> evar_map -> int -> constr -> Context.Rel.t * constr -val splay_lam_n : env -> evar_map -> int -> constr -> Context.Rel.t * constr +val splay_prod_n : env -> evar_map -> int -> constr -> rel_context * constr +val splay_lam_n : env -> evar_map -> int -> constr -> rel_context * constr val splay_prod_assum : - env -> evar_map -> constr -> Context.Rel.t * constr + env -> evar_map -> constr -> rel_context * constr type 'a miota_args = { mP : constr; (** the result type *) diff --git a/pretyping/retyping.ml b/pretyping/retyping.ml index a9529d560..bb1b2901e 100644 --- a/pretyping/retyping.ml +++ b/pretyping/retyping.ml @@ -50,14 +50,6 @@ let anomaly_on_error f x = try f x with RetypeError e -> anomaly ~label:"retyping" (print_retype_error e) -let local_assum (na, t) = - let inj = EConstr.Unsafe.to_constr in - LocalAssum (na, inj t) - -let local_def (na, b, t) = - let inj = EConstr.Unsafe.to_constr in - LocalDef (na, inj b, inj t) - let get_type_from_constraints env sigma t = if isEvar sigma (fst (decompose_app_vect sigma t)) then match @@ -84,13 +76,13 @@ let rec subst_type env sigma typ = function let sort_of_atomic_type env sigma ft args = let rec concl_of_arity env n ar args = match EConstr.kind sigma (whd_all env sigma ar), args with - | Prod (na, t, b), h::l -> concl_of_arity (push_rel (local_def (na, lift n h, t)) env) (n + 1) b l + | Prod (na, t, b), h::l -> concl_of_arity (push_rel (LocalDef (na, lift n h, t)) env) (n + 1) b l | Sort s, [] -> s | _ -> retype_error NotASort in concl_of_arity env 0 ft (Array.to_list args) let type_of_var env id = - try EConstr.of_constr (NamedDecl.get_type (lookup_named id env)) + try NamedDecl.get_type (lookup_named id env) with Not_found -> retype_error (BadVariable id) let decomp_sort env sigma t = @@ -105,7 +97,7 @@ let retype ?(polyprop=true) sigma = (try strip_outer_cast sigma (EConstr.of_constr (Evd.meta_ftype sigma n).Evd.rebus) with Not_found -> retype_error (BadMeta n)) | Rel n -> - let ty = EConstr.of_constr (RelDecl.get_type (lookup_rel n env)) in + let ty = RelDecl.get_type (lookup_rel n env) in lift n ty | Var id -> type_of_var env id | Const cst -> EConstr.of_constr (rename_type_of_constant env cst) @@ -128,9 +120,9 @@ let retype ?(polyprop=true) sigma = | Prod _ -> whd_beta sigma (applist (t, [c])) | _ -> t) | Lambda (name,c1,c2) -> - mkProd (name, c1, type_of (push_rel (local_assum (name,c1)) env) c2) + mkProd (name, c1, type_of (push_rel (LocalAssum (name,c1)) env) c2) | LetIn (name,b,c1,c2) -> - subst1 b (type_of (push_rel (local_def (name,b,c1)) env) c2) + subst1 b (type_of (push_rel (LocalDef (name,b,c1)) env) c2) | Fix ((_,i),(_,tys,_)) -> tys.(i) | CoFix (i,(_,tys,_)) -> tys.(i) | App(f,args) when is_template_polymorphic env sigma f -> @@ -153,7 +145,7 @@ let retype ?(polyprop=true) sigma = | Sort (Prop c) -> type1_sort | Sort (Type u) -> Type (Univ.super u) | Prod (name,t,c2) -> - (match (sort_of env t, sort_of (push_rel (local_assum (name,t)) env) c2) with + (match (sort_of env t, sort_of (push_rel (LocalAssum (name,t)) env) c2) with | _, (Prop Null as s) -> s | Prop _, (Prop Pos as s) -> s | Type _, (Prop Pos as s) when is_impredicative_set env -> s @@ -174,7 +166,7 @@ let retype ?(polyprop=true) sigma = | Sort (Prop c) -> InType | Sort (Type u) -> InType | Prod (name,t,c2) -> - let s2 = sort_family_of (push_rel (local_assum (name,t)) env) c2 in + let s2 = sort_family_of (push_rel (LocalAssum (name,t)) env) c2 in if not (is_impredicative_set env) && s2 == InSet && sort_family_of env t == InType then InType else s2 | App(f,args) when is_template_polymorphic env sigma f -> @@ -249,7 +241,7 @@ let sorts_of_context env evc ctxt = | [] -> env,[] | d :: ctxt -> let env,sorts = aux ctxt in - let s = get_sort_of env evc (EConstr.of_constr (RelDecl.get_type d)) in + let s = get_sort_of env evc (RelDecl.get_type d) in (push_rel d env,s::sorts) in snd (aux ctxt) diff --git a/pretyping/retyping.mli b/pretyping/retyping.mli index ce9e1635f..25129db1c 100644 --- a/pretyping/retyping.mli +++ b/pretyping/retyping.mli @@ -44,7 +44,7 @@ val type_of_global_reference_knowing_parameters : env -> evar_map -> constr -> val type_of_global_reference_knowing_conclusion : env -> evar_map -> constr -> types -> evar_map * types -val sorts_of_context : env -> evar_map -> Context.Rel.t -> sorts list +val sorts_of_context : env -> evar_map -> rel_context -> sorts list val expand_projection : env -> evar_map -> Names.projection -> constr -> constr list -> constr diff --git a/pretyping/tacred.ml b/pretyping/tacred.ml index 4abfc26fc..9f3f3c7e5 100644 --- a/pretyping/tacred.ml +++ b/pretyping/tacred.ml @@ -14,11 +14,11 @@ open Term open Libnames open Globnames open Termops +open Environ open EConstr open Vars open Find_subterm open Namegen -open Environ open CClosure open Reductionops open Cbv @@ -60,7 +60,7 @@ let is_evaluable env = function let value_of_evaluable_ref env evref u = match evref with | EvalConstRef con -> - (try constant_value_in env (con,u) + EConstr.of_constr (try constant_value_in env (con,u) with NotEvaluableConst IsProj -> raise (Invalid_argument "value_of_evaluable_ref")) | EvalVarRef id -> env |> lookup_named id |> NamedDecl.get_value |> Option.get @@ -115,9 +115,9 @@ let unsafe_reference_opt_value env sigma eval = | Declarations.Def c -> Some (EConstr.of_constr (Mod_subst.force_constr c)) | _ -> None) | EvalVar id -> - env |> lookup_named id |> NamedDecl.get_value |> Option.map EConstr.of_constr + env |> lookup_named id |> NamedDecl.get_value | EvalRel n -> - env |> lookup_rel n |> RelDecl.get_value |> Option.map (EConstr.of_constr %> lift n) + env |> lookup_rel n |> RelDecl.get_value |> Option.map (lift n) | EvalEvar ev -> match EConstr.kind sigma (mkEvar ev) with | Evar _ -> None @@ -127,9 +127,9 @@ let reference_opt_value env sigma eval u = match eval with | EvalConst cst -> Option.map EConstr.of_constr (constant_opt_value_in env (cst,u)) | EvalVar id -> - env |> lookup_named id |> NamedDecl.get_value |> Option.map EConstr.of_constr + env |> lookup_named id |> NamedDecl.get_value | EvalRel n -> - env |> lookup_rel n |> RelDecl.get_value |> Option.map (EConstr.of_constr %> lift n) + env |> lookup_rel n |> RelDecl.get_value |> Option.map (lift n) | EvalEvar ev -> match EConstr.kind sigma (mkEvar ev) with | Evar _ -> None @@ -146,11 +146,11 @@ let reference_value env sigma c u = (* One reuses the name of the function after reduction of the fixpoint *) type constant_evaluation = - | EliminationFix of int * int * (int * (int * EConstr.t) list * int) + | EliminationFix of int * int * (int * (int * constr) list * int) | EliminationMutualFix of int * evaluable_reference * ((int*evaluable_reference) option array * - (int * (int * EConstr.t) list * int)) + (int * (int * constr) list * int)) | EliminationCases of int | EliminationProj of int | NotAnElimination @@ -261,22 +261,13 @@ let invert_name labs l na0 env sigma ref = function [compute_consteval_mutual_fix] only one by one, until finding the last one before the Fix if the latter is mutually defined *) -let local_assum (na, t) = - let open Context.Rel.Declaration in - let inj = EConstr.Unsafe.to_constr in - LocalAssum (na, inj t) - -let local_def (na, b, t) = - let open Context.Rel.Declaration in - let inj = EConstr.Unsafe.to_constr in - LocalDef (na, inj b, inj t) - let compute_consteval_direct env sigma ref = let rec srec env n labs onlyproj c = let c',l = whd_betadeltazeta_stack env sigma c in match EConstr.kind sigma c' with | Lambda (id,t,g) when List.is_empty l && not onlyproj -> - srec (push_rel (local_assum (id,t)) env) (n+1) (t::labs) onlyproj g + let open Context.Rel.Declaration in + srec (push_rel (LocalAssum (id,t)) env) (n+1) (t::labs) onlyproj g | Fix fix when not onlyproj -> (try check_fix_reversibility sigma labs l fix with Elimconst -> NotAnElimination) @@ -295,7 +286,8 @@ let compute_consteval_mutual_fix env sigma ref = let nargs = List.length l in match EConstr.kind sigma c' with | Lambda (na,t,g) when List.is_empty l -> - srec (push_rel (local_assum (na,t)) env) (minarg+1) (t::labs) ref g + let open Context.Rel.Declaration in + srec (push_rel (LocalAssum (na,t)) env) (minarg+1) (t::labs) ref g | Fix ((lv,i),(names,_,_)) -> (* Last known constant wrapping Fix is ref = [labs](Fix l) *) (match compute_consteval_direct env sigma ref with @@ -386,7 +378,7 @@ let make_elim_fun (names,(nbfix,lv,n)) u largs = (* [f] is convertible to [Fix(recindices,bodynum),bodyvect)]: do so that the reduction uses this extra information *) -let dummy = Constr.mkProp +let dummy = mkProp let vfx = Id.of_string "_expanded_fix_" let vfun = Id.of_string "_eliminator_function_" let venv = let open Context.Named.Declaration in @@ -405,7 +397,7 @@ let substl_with_function subst sigma constr = match v.(i-k-1) with | (fx, Some (min, ref)) -> let sigma = Sigma.Unsafe.of_evar_map !evd in - let Sigma (evk, sigma, _) = Evarutil.new_pure_evar venv sigma (EConstr.of_constr dummy) in + let Sigma (evk, sigma, _) = Evarutil.new_pure_evar venv sigma dummy in let sigma = Sigma.to_evar_map sigma in evd := sigma; minargs := Evar.Map.add evk min !minargs; @@ -466,7 +458,7 @@ let substl_checking_arity env subst sigma c = in nf_fix body -type fix_reduction_result = NotReducible | Reduced of (EConstr.t * EConstr.t list) +type fix_reduction_result = NotReducible | Reduced of (constr * constr list) let reduce_fix whdfun sigma fix stack = match fix_recarg fix (Stack.append_app_list stack Stack.empty) with @@ -557,9 +549,9 @@ let match_eval_ref_value env sigma constr = | Const (sp, u) when is_evaluable env (EvalConstRef sp) -> Some (EConstr.of_constr (constant_value_in env (sp, u))) | Var id when is_evaluable env (EvalVarRef id) -> - env |> lookup_named id |> NamedDecl.get_value |> Option.map EConstr.of_constr + env |> lookup_named id |> NamedDecl.get_value | Rel n -> - env |> lookup_rel n |> RelDecl.get_value |> Option.map (EConstr.of_constr %> lift n) + env |> lookup_rel n |> RelDecl.get_value |> Option.map (lift n) | _ -> None let special_red_case env sigma whfun (ci, p, c, lf) = @@ -625,12 +617,12 @@ let whd_nothing_for_iota env sigma s = | Rel n -> let open Context.Rel.Declaration in (match lookup_rel n env with - | LocalDef (_,body,_) -> whrec (lift n (EConstr.of_constr body), stack) + | LocalDef (_,body,_) -> whrec (lift n body, stack) | _ -> s) | Var id -> let open Context.Named.Declaration in (match lookup_named id env with - | LocalDef (_,body,_) -> whrec (EConstr.of_constr body, stack) + | LocalDef (_,body,_) -> whrec (body, stack) | _ -> s) | Evar ev -> s | Meta ev -> @@ -834,7 +826,8 @@ let try_red_product env sigma c = | _ -> simpfun (mkApp (redrec env f, l))) | Cast (c,_,_) -> redrec env c | Prod (x,a,b) -> - mkProd (x, a, redrec (push_rel (local_assum (x, a)) env) b) + let open Context.Rel.Declaration in + mkProd (x, a, redrec (push_rel (LocalAssum (x, a)) env) b) | LetIn (x,a,b,t) -> redrec env (Vars.subst1 a t) | Case (ci,p,d,lf) -> simpfun (mkCase (ci,p,redrec env d,lf)) | Proj (p, c) -> @@ -1053,7 +1046,7 @@ let substlin env sigma evalref n (nowhere_except_in,locs) c = let maxocc = List.fold_right max locs 0 in let pos = ref n in assert (List.for_all (fun x -> x >= 0) locs); - let value u = EConstr.of_constr (value_of_evaluable_ref env evalref u) in + let value u = value_of_evaluable_ref env evalref u in let rec substrec () c = if nowhere_except_in && !pos > maxocc then c else @@ -1192,7 +1185,7 @@ let reduce_to_ind_gen allow_product env sigma t = | Prod (n,ty,t') -> let open Context.Rel.Declaration in if allow_product then - elimrec (push_rel (local_assum (n,ty)) env) t' ((local_assum (n,ty))::l) + elimrec (push_rel (LocalAssum (n,ty)) env) t' ((LocalAssum (n,ty))::l) else user_err (str"Not an inductive definition.") | _ -> @@ -1270,7 +1263,7 @@ let reduce_to_ref_gen allow_product env sigma ref t = | Prod (n,ty,t') -> if allow_product then let open Context.Rel.Declaration in - elimrec (push_rel (local_assum (n,t)) env) t' ((local_assum (n,ty))::l) + elimrec (push_rel (LocalAssum (n,t)) env) t' ((LocalAssum (n,ty))::l) else error_cannot_recognize ref | _ -> diff --git a/pretyping/typeclasses.ml b/pretyping/typeclasses.ml index 50ae66eb0..ce570ee12 100644 --- a/pretyping/typeclasses.ml +++ b/pretyping/typeclasses.ml @@ -300,6 +300,7 @@ let build_subclasses ~check env sigma glob pri = | Some (Backward, _) -> None | Some (Forward, pri') -> let proj = Option.get proj in + let rels = List.map (fun d -> Termops.map_rel_decl EConstr.Unsafe.to_constr d) rels in let body = it_mkLambda_or_LetIn (mkApp (mkConstU (proj,u), projargs)) rels in if check && check_instance env sigma (EConstr.of_constr body) then None else diff --git a/pretyping/typeclasses.mli b/pretyping/typeclasses.mli index e95aba695..0c30296d3 100644 --- a/pretyping/typeclasses.mli +++ b/pretyping/typeclasses.mli @@ -67,7 +67,7 @@ val dest_class_app : env -> evar_map -> EConstr.constr -> typeclass puniverses * val typeclass_univ_instance : typeclass puniverses -> typeclass puniverses (** Just return None if not a class *) -val class_of_constr : evar_map -> EConstr.constr -> (Context.Rel.t * (typeclass puniverses * constr list)) option +val class_of_constr : evar_map -> EConstr.constr -> (EConstr.rel_context * (typeclass puniverses * constr list)) option val instance_impl : instance -> global_reference diff --git a/pretyping/typing.ml b/pretyping/typing.ml index e6f1e46b6..bdd3663d1 100644 --- a/pretyping/typing.ml +++ b/pretyping/typing.ml @@ -12,9 +12,9 @@ open Pp open CErrors open Util open Term +open Environ open EConstr open Vars -open Environ open Reductionops open Inductive open Inductiveops @@ -23,14 +23,6 @@ open Arguments_renaming open Pretype_errors open Context.Rel.Declaration -let local_assum (na, t) = - let inj = EConstr.Unsafe.to_constr in - LocalAssum (na, inj t) - -let local_def (na, b, t) = - let inj = EConstr.Unsafe.to_constr in - LocalDef (na, inj b, inj t) - let push_rec_types pfix env = let (i, c, t) = pfix in let inj c = EConstr.Unsafe.to_constr c in @@ -101,14 +93,15 @@ let max_sort l = let e_is_correct_arity env evdref c pj ind specif params = let arsign = make_arity_signature env true (make_ind_family (ind,params)) in + let arsign = List.map (fun d -> Termops.map_rel_decl EConstr.of_constr d) arsign in let allowed_sorts = elim_sorts specif in let error () = Pretype_errors.error_elim_arity env !evdref ind allowed_sorts c pj None in let rec srec env pt ar = let pt' = whd_all env !evdref pt in match EConstr.kind !evdref pt', ar with | Prod (na1,a1,t), (LocalAssum (_,a1'))::ar' -> - if not (Evarconv.e_cumul env evdref a1 (EConstr.of_constr a1')) then error (); - srec (push_rel (local_assum (na1,a1)) env) t ar' + if not (Evarconv.e_cumul env evdref a1 a1') then error (); + srec (push_rel (LocalAssum (na1,a1)) env) t ar' | Sort s, [] -> if not (Sorts.List.mem (Sorts.family s) allowed_sorts) then error () @@ -326,14 +319,14 @@ let rec execute env evdref cstr = | Lambda (name,c1,c2) -> let j = execute env evdref c1 in let var = e_type_judgment env evdref j in - let env1 = push_rel (local_assum (name, var.utj_val)) env in + let env1 = push_rel (LocalAssum (name, var.utj_val)) env in let j' = execute env1 evdref c2 in judge_of_abstraction env1 name var j' | Prod (name,c1,c2) -> let j = execute env evdref c1 in let varj = e_type_judgment env evdref j in - let env1 = push_rel (local_assum (name, varj.utj_val)) env in + let env1 = push_rel (LocalAssum (name, varj.utj_val)) env in let j' = execute env1 evdref c2 in let varj' = e_type_judgment env1 evdref j' in judge_of_product env name varj varj' @@ -343,7 +336,7 @@ let rec execute env evdref cstr = let j2 = execute env evdref c2 in let j2 = e_type_judgment env evdref j2 in let _ = e_judge_of_cast env evdref j1 DEFAULTcast j2 in - let env1 = push_rel (local_def (name, j1.uj_val, j2.utj_val)) env in + let env1 = push_rel (LocalDef (name, j1.uj_val, j2.utj_val)) env in let j3 = execute env1 evdref c3 in judge_of_letin env name j1 j2 j3 diff --git a/pretyping/unification.ml b/pretyping/unification.ml index 04cc4253e..0d6dcffc1 100644 --- a/pretyping/unification.ml +++ b/pretyping/unification.ml @@ -14,10 +14,10 @@ open Util open Names open Term open Termops +open Environ open EConstr open Vars open Namegen -open Environ open Evd open Reduction open Reductionops @@ -91,7 +91,6 @@ let abstract_scheme env evd c l lname_typ = (fun (t,evd) (locc,a) decl -> let na = RelDecl.get_name decl in let ta = RelDecl.get_type decl in - let ta = EConstr.of_constr ta in let na = match EConstr.kind evd a with Var id -> Name id | _ -> na in (* [occur_meta ta] test removed for support of eelim/ecase but consequences are unclear... @@ -1627,6 +1626,7 @@ let make_abstraction_core name (test,out) env sigma c ty occs check_occs concl = let likefirst = clause_with_generic_occurrences occs in let mkvarid () = EConstr.mkVar id in let compute_dependency _ d (sign,depdecls) = + let d = map_named_decl EConstr.of_constr d in let hyp = NamedDecl.get_id d in match occurrences_of_hyp hyp occs with | NoOccurrences, InHyp -> @@ -1634,7 +1634,7 @@ let make_abstraction_core name (test,out) env sigma c ty occs check_occs concl = | AllOccurrences, InHyp as occ -> let occ = if likefirst then LikeFirst else AtOccs occ in let newdecl = replace_term_occ_decl_modulo sigma occ test mkvarid d in - if Context.Named.Declaration.equal Constr.equal d newdecl + if Context.Named.Declaration.equal (EConstr.eq_constr sigma) d newdecl && not (indirectly_dependent sigma c d depdecls) then if check_occs && not (in_every_hyp occs) @@ -1688,7 +1688,7 @@ type abstraction_request = type 'r abstraction_result = Names.Id.t * named_context_val * - Context.Named.Declaration.t list * Names.Id.t option * + named_declaration list * Names.Id.t option * types * (constr, 'r) Sigma.sigma option let make_abstraction env evd ccl abs = diff --git a/pretyping/unification.mli b/pretyping/unification.mli index 41dcb8ed3..6760283d2 100644 --- a/pretyping/unification.mli +++ b/pretyping/unification.mli @@ -82,7 +82,7 @@ val finish_evar_resolution : ?flags:Pretyping.inference_flags -> type 'r abstraction_result = Names.Id.t * named_context_val * - Context.Named.Declaration.t list * Names.Id.t option * + named_declaration list * Names.Id.t option * types * (constr, 'r) Sigma.sigma option val make_abstraction : env -> 'r Sigma.t -> constr -> diff --git a/printing/prettyp.ml b/printing/prettyp.ml index 93970512d..64c7c0ba5 100644 --- a/printing/prettyp.ml +++ b/printing/prettyp.ml @@ -71,11 +71,11 @@ let print_basename sp = pr_global (ConstRef sp) let print_ref reduce ref = let typ = Global.type_of_global_unsafe ref in + let typ = EConstr.of_constr typ in let typ = if reduce then - let ctx,ccl = Reductionops.splay_prod_assum (Global.env()) Evd.empty (EConstr.of_constr typ) in - let ccl = EConstr.Unsafe.to_constr ccl - in Term.it_mkProd_or_LetIn ccl ctx + let ctx,ccl = Reductionops.splay_prod_assum (Global.env()) Evd.empty typ + in EConstr.it_mkProd_or_LetIn ccl ctx else typ in let univs = Global.universes_of_global ref in let env = Global.env () in @@ -85,7 +85,7 @@ let print_ref reduce ref = if Global.is_polymorphic ref then Printer.pr_universe_instance sigma univs else mt () in - hov 0 (pr_global ref ++ inst ++ str " :" ++ spc () ++ pr_ltype_env env sigma typ ++ + hov 0 (pr_global ref ++ inst ++ str " :" ++ spc () ++ pr_letype_env env sigma typ ++ Printer.pr_universe_ctx sigma univs) (********************************) diff --git a/proofs/goal.ml b/proofs/goal.ml index 410e738de..9046f4534 100644 --- a/proofs/goal.ml +++ b/proofs/goal.ml @@ -155,6 +155,7 @@ module V82 = struct with Not_found -> true in Environ.fold_named_context_reverse (fun t decl -> if is_proof_var decl then + let decl = Termops.map_named_decl EConstr.of_constr decl in mkNamedProd_or_LetIn decl t else t diff --git a/proofs/logic.ml b/proofs/logic.ml index 59ce8ffa6..a31cadd88 100644 --- a/proofs/logic.ml +++ b/proofs/logic.ml @@ -174,12 +174,14 @@ let reorder_context env sigma sign ord = step ord ords sign mt_q [] let reorder_val_context env sigma sign ord = + let open EConstr in val_of_named_context (reorder_context env sigma (named_context_of_val sign) ord) let check_decl_position env sigma sign d = + let open EConstr in let x = NamedDecl.get_id d in let needed = global_vars_set_of_decl env sigma d in let deps = dependency_closure env sigma (named_context_of_val sign) needed in @@ -266,6 +268,7 @@ let move_hyp sigma toleft (left,declfrom,right) hto = else moverec first' middle' right in + let open EConstr in if toleft then let right = List.fold_right push_named_context_val right empty_named_context_val in @@ -279,6 +282,7 @@ let move_hyp sigma toleft (left,declfrom,right) hto = right left let move_hyp_in_named_context sigma hfrom hto sign = + let open EConstr in let (left,right,declfrom,toleft) = split_sign hfrom hto (named_context_of_val sign) in move_hyp sigma toleft (left,declfrom,right) hto @@ -507,7 +511,7 @@ and mk_casegoals sigma goal goalacc p c = let convert_hyp check sign sigma d = let id = NamedDecl.get_id d in - let b = Option.map EConstr.of_constr (NamedDecl.get_value d) in + let b = NamedDecl.get_value d in let env = Global.env() in let reorder = ref [] in let sign' = @@ -515,14 +519,14 @@ let convert_hyp check sign sigma d = (fun _ d' _ -> let c = Option.map EConstr.of_constr (NamedDecl.get_value d') in let env = Global.env_of_context sign in - if check && not (is_conv env sigma (EConstr.of_constr (NamedDecl.get_type d)) (EConstr.of_constr (NamedDecl.get_type d'))) then + if check && not (is_conv env sigma (NamedDecl.get_type d) (EConstr.of_constr (NamedDecl.get_type d'))) then user_err ~hdr:"Logic.convert_hyp" (str "Incorrect change of the type of " ++ pr_id id ++ str "."); if check && not (Option.equal (is_conv env sigma) b c) then user_err ~hdr:"Logic.convert_hyp" (str "Incorrect change of the body of "++ pr_id id ++ str "."); if check then reorder := check_decl_position env sigma sign d; - d) in + map_named_decl EConstr.Unsafe.to_constr d) in reorder_val_context env sigma sign' !reorder @@ -535,15 +539,16 @@ let prim_refiner r sigma goal = let env = Goal.V82.env sigma goal in let sign = Goal.V82.hyps sigma goal in let cl = Goal.V82.concl sigma goal in - let cl = EConstr.Unsafe.to_constr cl in let mk_goal hyps concl = Goal.V82.mk_goal sigma hyps concl (Goal.V82.extra sigma goal) in + let open EConstr in match r with (* Logical rules *) | Cut (b,replace,id,t) -> (* if !check && not (Retyping.get_sort_of env sigma t) then*) - let (sg1,ev1,sigma) = mk_goal sign (nf_betaiota sigma (EConstr.of_constr t)) in + let t = EConstr.of_constr t in + let (sg1,ev1,sigma) = mk_goal sign (nf_betaiota sigma t) in let sign,t,cl,sigma = if replace then let nexthyp = get_hyp_after id (named_context_of_val sign) in @@ -556,14 +561,14 @@ let prim_refiner r sigma goal = user_err ~hdr:"Logic.prim_refiner" (str "Variable " ++ pr_id id ++ str " is already declared."); push_named_context_val (LocalAssum (id,t)) sign,t,cl,sigma) in - let t = EConstr.of_constr t in let (sg2,ev2,sigma) = - Goal.V82.mk_goal sigma sign (EConstr.of_constr cl) (Goal.V82.extra sigma goal) in - let oterm = EConstr.mkLetIn (Name id, ev1, t, EConstr.Vars.subst_var id ev2) in + Goal.V82.mk_goal sigma sign cl (Goal.V82.extra sigma goal) in + let oterm = mkLetIn (Name id, ev1, t, EConstr.Vars.subst_var id ev2) in let sigma = Goal.V82.partial_solution_to sigma goal sg2 oterm in if b then ([sg1;sg2],sigma) else ([sg2;sg1],sigma) | Refine c -> + let cl = EConstr.Unsafe.to_constr cl in check_meta_variables c; let (sgl,cl',sigma,oterm) = mk_refgoals sigma goal [] cl c in let sgl = List.rev sgl in diff --git a/proofs/logic.mli b/proofs/logic.mli index 5fe28ac76..f98248e4d 100644 --- a/proofs/logic.mli +++ b/proofs/logic.mli @@ -55,7 +55,7 @@ exception RefinerError of refiner_error val catchable_exception : exn -> bool val convert_hyp : bool -> Environ.named_context_val -> evar_map -> - Context.Named.Declaration.t -> Environ.named_context_val + EConstr.named_declaration -> Environ.named_context_val val move_hyp_in_named_context : Evd.evar_map -> Id.t -> Id.t Misctypes.move_location -> Environ.named_context_val -> Environ.named_context_val diff --git a/proofs/refiner.ml b/proofs/refiner.ml index 9a0b56b84..5c7659ac0 100644 --- a/proofs/refiner.ml +++ b/proofs/refiner.ml @@ -22,7 +22,7 @@ let project x = x.sigma (* Getting env *) let pf_env gls = Global.env_of_context (Goal.V82.hyps (project gls) (sig_it gls)) -let pf_hyps gls = named_context_of_val (Goal.V82.hyps (project gls) (sig_it gls)) +let pf_hyps gls = EConstr.named_context_of_val (Goal.V82.hyps (project gls) (sig_it gls)) let refiner pr goal_sigma = let (sgl,sigma') = prim_refiner pr goal_sigma.sigma goal_sigma.it in @@ -198,10 +198,10 @@ let tclNOTSAMEGOAL (tac : tactic) goal = destruct), this is not detected by this tactical. *) let tclSHOWHYPS (tac : tactic) (goal: Goal.goal Evd.sigma) :Proof_type.goal list Evd.sigma = - let oldhyps:Context.Named.t = pf_hyps goal in + let oldhyps = pf_hyps goal in let rslt:Proof_type.goal list Evd.sigma = tac goal in let { it = gls; sigma = sigma; } = rslt in - let hyps:Context.Named.t list = + let hyps = List.map (fun gl -> pf_hyps { it = gl; sigma=sigma; }) gls in let cmp d1 d2 = Names.Id.equal (NamedDecl.get_id d1) (NamedDecl.get_id d2) in let newhyps = diff --git a/proofs/refiner.mli b/proofs/refiner.mli index 6dcafb77a..56f5facf8 100644 --- a/proofs/refiner.mli +++ b/proofs/refiner.mli @@ -10,6 +10,7 @@ open Evd open Proof_type +open EConstr (** The refiner (handles primitive rules and high-level tactics). *) @@ -17,7 +18,7 @@ val sig_it : 'a sigma -> 'a val project : 'a sigma -> evar_map val pf_env : goal sigma -> Environ.env -val pf_hyps : goal sigma -> Context.Named.t +val pf_hyps : goal sigma -> named_context val unpackage : 'a sigma -> evar_map ref * 'a val repackage : evar_map ref -> 'a -> 'a sigma diff --git a/proofs/tacmach.ml b/proofs/tacmach.ml index 3ed262d6e..4b027a127 100644 --- a/proofs/tacmach.ml +++ b/proofs/tacmach.ml @@ -64,7 +64,7 @@ let pf_get_hyp gls id = raise (RefinerError (NoSuchHyp id)) let pf_get_hyp_typ gls id = - id |> pf_get_hyp gls |> NamedDecl.get_type |> EConstr.of_constr + id |> pf_get_hyp gls |> NamedDecl.get_type let pf_ids_of_hyps gls = ids_of_named_context (pf_hyps gls) @@ -199,13 +199,13 @@ module New = struct let pf_get_hyp id gl = let hyps = Proofview.Goal.env gl in let sign = - try Environ.lookup_named id hyps + try EConstr.lookup_named id hyps with Not_found -> raise (RefinerError (NoSuchHyp id)) in sign let pf_get_hyp_typ id gl = - pf_get_hyp id gl |> NamedDecl.get_type |> EConstr.of_constr + pf_get_hyp id gl |> NamedDecl.get_type let pf_hyps_types gl = let env = Proofview.Goal.env gl in diff --git a/proofs/tacmach.mli b/proofs/tacmach.mli index efc3dbf55..3b23a6ab4 100644 --- a/proofs/tacmach.mli +++ b/proofs/tacmach.mli @@ -34,18 +34,18 @@ val apply_sig_tac : val pf_concl : goal sigma -> types val pf_env : goal sigma -> env -val pf_hyps : goal sigma -> Context.Named.t +val pf_hyps : goal sigma -> named_context (*i val pf_untyped_hyps : goal sigma -> (Id.t * constr) list i*) val pf_hyps_types : goal sigma -> (Id.t * types) list val pf_nth_hyp_id : goal sigma -> int -> Id.t -val pf_last_hyp : goal sigma -> Context.Named.Declaration.t +val pf_last_hyp : goal sigma -> named_declaration val pf_ids_of_hyps : goal sigma -> Id.t list val pf_global : goal sigma -> Id.t -> constr val pf_unsafe_type_of : goal sigma -> constr -> types val pf_type_of : goal sigma -> constr -> evar_map * types val pf_hnf_type_of : goal sigma -> constr -> types -val pf_get_hyp : goal sigma -> Id.t -> Context.Named.Declaration.t +val pf_get_hyp : goal sigma -> Id.t -> named_declaration val pf_get_hyp_typ : goal sigma -> Id.t -> types val pf_get_new_id : Id.t -> goal sigma -> Id.t @@ -117,9 +117,9 @@ module New : sig val pf_ids_of_hyps : ('a, 'r) Proofview.Goal.t -> identifier list val pf_hyps_types : ('a, 'r) Proofview.Goal.t -> (identifier * types) list - val pf_get_hyp : identifier -> ([ `NF ], 'r) Proofview.Goal.t -> Context.Named.Declaration.t + val pf_get_hyp : identifier -> ([ `NF ], 'r) Proofview.Goal.t -> named_declaration val pf_get_hyp_typ : identifier -> ([ `NF ], 'r) Proofview.Goal.t -> types - val pf_last_hyp : ([ `NF ], 'r) Proofview.Goal.t -> Context.Named.Declaration.t + val pf_last_hyp : ([ `NF ], 'r) Proofview.Goal.t -> named_declaration val pf_nf_concl : ([ `LZ ], 'r) Proofview.Goal.t -> types val pf_reduce_to_quantified_ind : ('a, 'r) Proofview.Goal.t -> types -> pinductive * types diff --git a/tactics/auto.ml b/tactics/auto.ml index 17a488ddb..b548f8b92 100644 --- a/tactics/auto.ml +++ b/tactics/auto.ml @@ -321,7 +321,7 @@ let rec trivial_fail_db dbg mod_delta db_list local_db = ( Proofview.Goal.enter { enter = begin fun gl -> let sigma = Tacmach.New.project gl in let env = Proofview.Goal.env gl in - let nf c = EConstr.Unsafe.to_constr (Evarutil.nf_evar sigma (EConstr.of_constr c)) in + let nf c = Evarutil.nf_evar sigma c in let decl = Tacmach.New.pf_last_hyp (Proofview.Goal.assume gl) in let hyp = Context.Named.Declaration.map_constr nf decl in let hintl = make_resolve_hyp env sigma hyp diff --git a/tactics/class_tactics.ml b/tactics/class_tactics.ml index ef67d28f9..55fda1c7d 100644 --- a/tactics/class_tactics.ml +++ b/tactics/class_tactics.ml @@ -521,14 +521,14 @@ let evars_to_goals p evm = (** Making local hints *) let make_resolve_hyp env sigma st flags only_classes pri decl = let id = NamedDecl.get_id decl in - let cty = Evarutil.nf_evar sigma (EConstr.of_constr (NamedDecl.get_type decl)) in + let cty = Evarutil.nf_evar sigma (NamedDecl.get_type decl) in let rec iscl env ty = let ctx, ar = decompose_prod_assum sigma ty in match EConstr.kind sigma (fst (decompose_app sigma ar)) with | Const (c,_) -> is_class (ConstRef c) | Ind (i,_) -> is_class (IndRef i) | _ -> - let env' = Environ.push_rel_context ctx env in + let env' = push_rel_context ctx env in let ty' = Reductionops.whd_all env' sigma ar in if not (EConstr.eq_constr sigma ty' ar) then iscl env' ty' else false @@ -562,7 +562,7 @@ let make_hints g st only_classes sign = let consider = try let t = hyp |> NamedDecl.get_id |> Global.lookup_named |> NamedDecl.get_type in (* Section variable, reindex only if the type changed *) - not (Term.eq_constr t (NamedDecl.get_type hyp)) + not (EConstr.eq_constr (project g) (EConstr.of_constr t) (NamedDecl.get_type hyp)) with Not_found -> true in if consider then @@ -617,7 +617,7 @@ module V85 = struct then cached_hints else - let hints = make_hints g st only_classes (Environ.named_context_of_val sign) + let hints = make_hints g st only_classes (EConstr.named_context_of_val sign) in cache := (only_classes, sign, hints); hints @@ -634,7 +634,7 @@ module V85 = struct let gls' = List.map (fun g' -> let env = Goal.V82.env s g' in - let context = Environ.named_context_of_val (Goal.V82.hyps s g') in + let context = EConstr.named_context_of_val (Goal.V82.hyps s g') in let hint = make_resolve_hyp env s (Hint_db.transparent_state info.hints) (true,false,false) info.only_classes None (List.hd context) in let ldb = Hint_db.add_list env s hint info.hints in @@ -937,9 +937,10 @@ module Search = struct let sign = Goal.hyps g in let (dir, onlyc, sign', cached_hints) = !autogoal_cache in let cwd = Lib.cwd () in + let eq c1 c2 = EConstr.eq_constr (project g) c1 c2 in if DirPath.equal cwd dir && (onlyc == only_classes) && - Context.Named.equal Constr.equal sign sign' && + Context.Named.equal eq sign sign' && Hint_db.transparent_state cached_hints == st then cached_hints else @@ -1033,8 +1034,9 @@ module Search = struct Feedback.msg_debug (pr_depth (succ j :: i :: info.search_depth) ++ str" : " ++ pr_ev s' (Proofview.Goal.goal gl')); + let eq c1 c2 = EConstr.eq_constr s' c1 c2 in let hints' = - if b && not (Context.Named.equal Constr.equal (Goal.hyps gl') (Goal.hyps gl)) + if b && not (Context.Named.equal eq (Goal.hyps gl') (Goal.hyps gl)) then let st = Hint_db.transparent_state info.search_hints in make_autogoal_hints info.search_only_classes ~st gl' diff --git a/tactics/contradiction.ml b/tactics/contradiction.ml index 7173fb4fd..0e28aa980 100644 --- a/tactics/contradiction.ml +++ b/tactics/contradiction.ml @@ -51,7 +51,7 @@ let use_negated_unit_or_eq_type () = Flags.version_strictly_greater Flags.V8_5 let filter_hyp f tac = let rec seek = function | [] -> Proofview.tclZERO Not_found - | d::rest when f (EConstr.of_constr (NamedDecl.get_type d)) -> tac (NamedDecl.get_id d) + | d::rest when f (NamedDecl.get_type d) -> tac (NamedDecl.get_id d) | _::rest -> seek rest in Proofview.Goal.enter { enter = begin fun gl -> let hyps = Proofview.Goal.hyps (Proofview.Goal.assume gl) in @@ -66,7 +66,7 @@ let contradiction_context = | [] -> Tacticals.New.tclZEROMSG (Pp.str"No such contradiction") | d :: rest -> let id = NamedDecl.get_id d in - let typ = nf_evar sigma (EConstr.of_constr (NamedDecl.get_type d)) in + let typ = nf_evar sigma (NamedDecl.get_type d) in let typ = whd_all env sigma typ in if is_empty_type sigma typ then simplest_elim (mkVar id) diff --git a/tactics/equality.ml b/tactics/equality.ml index 072da995d..122b64777 100644 --- a/tactics/equality.ml +++ b/tactics/equality.ml @@ -13,12 +13,12 @@ open Names open Nameops open Term open Termops +open Environ open EConstr open Vars open Namegen open Inductive open Inductiveops -open Environ open Libnames open Globnames open Reductionops @@ -47,10 +47,6 @@ open Context.Named.Declaration module NamedDecl = Context.Named.Declaration -let nlocal_assum (na, t) = - let inj = EConstr.Unsafe.to_constr in - NamedDecl.LocalAssum (na, inj t) - (* Options *) let discriminate_introduction = ref true @@ -333,7 +329,7 @@ let jmeq_same_dom gl = function | None -> true (* already checked in Hipattern.find_eq_data_decompose *) | Some t -> let rels, t = decompose_prod_assum (project gl) t in - let env = Environ.push_rel_context rels (Proofview.Goal.env gl) in + let env = push_rel_context rels (Proofview.Goal.env gl) in match decompose_app (project gl) t with | _, [dom1; _; dom2;_] -> is_conv env (Tacmach.New.project gl) dom1 dom2 | _ -> false @@ -857,16 +853,19 @@ let descend_then env sigma head dirn = let (mib,mip) = lookup_mind_specif env ind in let cstr = get_constructors env indf in let dirn_nlams = cstr.(dirn-1).cs_nargs in - let dirn_env = push_rel_context cstr.(dirn-1).cs_args env in + let dirn_env = Environ.push_rel_context cstr.(dirn-1).cs_args env in (dirn_nlams, dirn_env, (fun dirnval (dfltval,resty) -> let deparsign = make_arity_signature env true indf in + let deparsign = List.map (fun d -> map_rel_decl EConstr.of_constr d) deparsign in let p = it_mkLambda_or_LetIn (lift (mip.mind_nrealargs+1) resty) deparsign in let build_branch i = let result = if Int.equal i dirn then dirnval else dfltval in - it_mkLambda_or_LetIn result (name_context env cstr.(i-1).cs_args) in + let args = name_context env cstr.(i-1).cs_args in + let args = List.map (fun d -> map_rel_decl EConstr.of_constr d) args in + it_mkLambda_or_LetIn result args in let brl = List.map build_branch (List.interval 1 (Array.length mip.mind_consnames)) in @@ -907,11 +906,13 @@ let build_selector env sigma dirn c ind special default = let typ = Retyping.get_type_of env sigma default in let (mib,mip) = lookup_mind_specif env ind in let deparsign = make_arity_signature env true indf in + let deparsign = List.map (fun d -> map_rel_decl EConstr.of_constr d) deparsign in let p = it_mkLambda_or_LetIn typ deparsign in let cstrs = get_constructors env indf in let build_branch i = let endpt = if Int.equal i dirn then special else default in - it_mkLambda_or_LetIn endpt cstrs.(i-1).cs_args in + let args = List.map (fun d -> map_rel_decl EConstr.of_constr d) cstrs.(i-1).cs_args in + it_mkLambda_or_LetIn endpt args in let brl = List.map build_branch(List.interval 1 (Array.length mip.mind_consnames)) in let ci = make_case_info env ind RegularStyle in @@ -995,7 +996,7 @@ let apply_on_clause (f,t) clause = let discr_positions env sigma (lbeq,eqn,(t,t1,t2)) eq_clause cpath dirn = let e = next_ident_away eq_baseid (ids_of_context env) in - let e_env = push_named (nlocal_assum (e, t)) env in + let e_env = push_named (Context.Named.Declaration.LocalAssum (e,t)) env in let discriminator = build_discriminator e_env sigma dirn (mkVar e) cpath in let sigma,(pf, absurd_term), eff = @@ -1372,7 +1373,7 @@ let simplify_args env sigma t = let inject_at_positions env sigma l2r (eq,_,(t,t1,t2)) eq_clause posns tac = let e = next_ident_away eq_baseid (ids_of_context env) in - let e_env = push_named (nlocal_assum (e,t)) env in + let e_env = push_named (LocalAssum (e,t)) env in let evdref = ref sigma in let filter (cpath, t1', t2') = try @@ -1696,7 +1697,7 @@ let is_eq_x gl x d = | Var id' -> Id.equal id id' | _ -> false in - let c = pf_nf_evar gl (EConstr.of_constr (NamedDecl.get_type d)) in + let c = pf_nf_evar gl (NamedDecl.get_type d) in let (_,lhs,rhs) = pi3 (find_eq_data_decompose gl c) in if (is_var x lhs) && not (local_occur_var (project gl) x rhs) then raise (FoundHyp (id,rhs,true)); if (is_var x rhs) && not (local_occur_var (project gl) x lhs) then raise (FoundHyp (id,lhs,false)) @@ -1793,7 +1794,7 @@ let subst_all ?(flags=default_subst_tactic_flags ()) () = let find_eq_data_decompose = find_eq_data_decompose gl in let select_equation_name decl = try - let lbeq,u,(_,x,y) = find_eq_data_decompose (EConstr.of_constr (NamedDecl.get_type decl)) in + let lbeq,u,(_,x,y) = find_eq_data_decompose (NamedDecl.get_type decl) in let eq = Universes.constr_of_global_univ (lbeq.eq,u) in if flags.only_leibniz then restrict_to_eq_and_identity eq; match EConstr.kind sigma x, EConstr.kind sigma y with @@ -1817,7 +1818,6 @@ let subst_all ?(flags=default_subst_tactic_flags ()) () = let env = Proofview.Goal.env gl in let find_eq_data_decompose = find_eq_data_decompose gl in let c = pf_get_hyp hyp gl |> NamedDecl.get_type in - let c = EConstr.of_constr c in let _,_,(_,x,y) = find_eq_data_decompose c in (* J.F.: added to prevent failure on goal containing x=x as an hyp *) if EConstr.eq_constr sigma x y then Proofview.tclUNIT () else @@ -1890,7 +1890,7 @@ let rewrite_assumption_cond cond_eq_term cl = let id = NamedDecl.get_id hyp in begin try - let dir = cond_eq_term (EConstr.of_constr (NamedDecl.get_type hyp)) gl in + let dir = cond_eq_term (NamedDecl.get_type hyp) gl in general_rewrite_clause dir false (mkVar id,NoBindings) cl with | Failure _ | UserError _ -> arec rest gl end diff --git a/tactics/hints.ml b/tactics/hints.ml index ef97b0b33..ffd19ac6e 100644 --- a/tactics/hints.ml +++ b/tactics/hints.ml @@ -875,7 +875,7 @@ let make_resolve_hyp env sigma decl = try [make_apply_entry env sigma (true, true, false) None false ~name:(PathHints [VarRef hname]) - (c, EConstr.of_constr (NamedDecl.get_type decl), Univ.ContextSet.empty)] + (c, NamedDecl.get_type decl, Univ.ContextSet.empty)] with | Failure _ -> [] | e when Logic.catchable_exception e -> anomaly (Pp.str "make_resolve_hyp") @@ -1335,7 +1335,7 @@ let make_local_hint_db env sigma ts eapply lems = (Sigma.to_evar_map sigma, c) in let lems = List.map map lems in - let sign = Environ.named_context env in + let sign = EConstr.named_context env in let ts = match ts with | None -> Hint_db.transparent_state (searchtable_map "core") | Some ts -> ts diff --git a/tactics/hints.mli b/tactics/hints.mli index 344827e03..0d6dd434e 100644 --- a/tactics/hints.mli +++ b/tactics/hints.mli @@ -29,7 +29,7 @@ val decompose_app_bound : evar_map -> constr -> global_reference * constr array type debug = Debug | Info | Off -val secvars_of_hyps : Context.Named.t -> Id.Pred.t +val secvars_of_hyps : ('c, 't) Context.Named.pt -> Id.Pred.t (** Pre-created hint databases *) @@ -209,7 +209,7 @@ val make_resolves : If the hyp cannot be used as a Hint, the empty list is returned. *) val make_resolve_hyp : - env -> evar_map -> Context.Named.Declaration.t -> hint_entry list + env -> evar_map -> named_declaration -> hint_entry list (** [make_extern pri pattern tactic_expr] *) diff --git a/tactics/hipattern.ml b/tactics/hipattern.ml index 607d6d2a9..8e4654c02 100644 --- a/tactics/hipattern.ml +++ b/tactics/hipattern.ml @@ -107,8 +107,8 @@ let match_with_one_constructor sigma style onlybinary allow_rec t = List.for_all (fun decl -> let c = RelDecl.get_type decl in is_local_assum decl && - Term.isRel c && - Int.equal (Term.destRel c) mib.mind_nparams) ctx + isRel sigma c && + Int.equal (destRel sigma c) mib.mind_nparams) ctx then Some (hdapp,args) else None @@ -117,7 +117,6 @@ let match_with_one_constructor sigma style onlybinary allow_rec t = let cargs = List.map RelDecl.get_type (prod_assum sigma ctyp) in if not (is_lax_conjunction style) || has_nodep_prod sigma ctyp then (* Record or non strict conjunction *) - let cargs = List.map EConstr.of_constr cargs in Some (hdapp,List.rev cargs) else None diff --git a/tactics/inv.ml b/tactics/inv.ml index 426749a75..ecb8eedac 100644 --- a/tactics/inv.ml +++ b/tactics/inv.ml @@ -13,10 +13,10 @@ open Names open Nameops open Term open Termops +open Environ open EConstr open Vars open Namegen -open Environ open Inductiveops open Printer open Retyping @@ -75,6 +75,7 @@ let make_inv_predicate env evd indf realargs id status concl = | NoDep -> (* We push the arity and leave concl unchanged *) let hyps_arity,_ = get_arity env indf in + let hyps_arity = List.map (fun d -> map_rel_decl EConstr.of_constr d) hyps_arity in (hyps_arity,concl) | Dep dflt_concl -> if not (occur_var env !evd id concl) then @@ -132,6 +133,10 @@ let make_inv_predicate env evd indf realargs id status concl = build_concl eqns args (succ n) restlist in let (newconcl, args) = build_concl [] [] 0 realargs in + let name_context env ctx = + let map f c = List.map (fun d -> Termops.map_rel_decl f d) c in + map EConstr.of_constr (name_context env (map EConstr.Unsafe.to_constr ctx)) + in let predicate = it_mkLambda_or_LetIn newconcl (name_context env hyps) in let _ = Evarutil.evd_comb1 (Typing.type_of env) evd predicate in (* OK - this predicate should now be usable by res_elimination_then to diff --git a/tactics/leminv.ml b/tactics/leminv.ml index a05b4fbf3..d7c396179 100644 --- a/tactics/leminv.ml +++ b/tactics/leminv.ml @@ -12,6 +12,7 @@ open Util open Names open Term open Termops +open Environ open EConstr open Vars open Namegen @@ -20,7 +21,6 @@ open Printer open Reductionops open Entries open Inductiveops -open Environ open Tacmach.New open Clenv open Declare @@ -32,14 +32,6 @@ open Context.Named.Declaration module NamedDecl = Context.Named.Declaration -let nlocal_assum (na, t) = - let inj = EConstr.Unsafe.to_constr in - NamedDecl.LocalAssum (na, inj t) - -let nlocal_def (na, b, t) = - let inj = EConstr.Unsafe.to_constr in - NamedDecl.LocalDef (na, inj b, inj t) - let no_inductive_inconstr env sigma constr = (str "Cannot recognize an inductive predicate in " ++ pr_leconstr_env env sigma constr ++ @@ -129,11 +121,11 @@ let rec add_prods_sign env sigma t = | Prod (na,c1,b) -> let id = id_of_name_using_hdchar env (EConstr.Unsafe.to_constr t) na in let b'= subst1 (mkVar id) b in - add_prods_sign (push_named (nlocal_assum (id,c1)) env) sigma b' + add_prods_sign (push_named (LocalAssum (id,c1)) env) sigma b' | LetIn (na,c1,t1,b) -> let id = id_of_name_using_hdchar env (EConstr.Unsafe.to_constr t) na in let b'= subst1 (mkVar id) b in - add_prods_sign (push_named (nlocal_def (id,c1,t1)) env) sigma b' + add_prods_sign (push_named (LocalDef (id,c1,t1)) env) sigma b' | _ -> (env,t) (* [dep_option] indicates whether the inversion lemma is dependent or not. @@ -168,6 +160,7 @@ let compute_first_inversion_scheme env sigma ind sort dep_option = let revargs,ownsign = fold_named_context (fun env d (revargs,hyps) -> + let d = map_named_decl EConstr.of_constr d in let id = NamedDecl.get_id d in if Id.List.mem id ivars then ((mkVar id)::revargs, Context.Named.add d hyps) @@ -180,7 +173,7 @@ let compute_first_inversion_scheme env sigma ind sort dep_option = (pty,goal) in let npty = nf_all env sigma pty in - let extenv = push_named (nlocal_assum (p,npty)) env in + let extenv = push_named (LocalAssum (p,npty)) env in extenv, goal (* [inversion_scheme sign I] @@ -218,6 +211,7 @@ let inversion_scheme env sigma t sort dep_option inv_op = let ownSign = ref begin fold_named_context (fun env d sign -> + let d = map_named_decl EConstr.of_constr d in if mem_named_context_val (NamedDecl.get_id d) global_named_context then sign else Context.Named.add d sign) invEnv ~init:Context.Named.empty @@ -231,7 +225,7 @@ let inversion_scheme env sigma t sort dep_option inv_op = let h = next_ident_away (Id.of_string "H") !avoid in let ty,inst = Evarutil.generalize_evar_over_rels sigma (e,args) in avoid := h::!avoid; - ownSign := Context.Named.add (nlocal_assum (h,ty)) !ownSign; + ownSign := Context.Named.add (LocalAssum (h,ty)) !ownSign; applist (mkVar h, inst) | _ -> EConstr.map sigma fill_holes c in diff --git a/tactics/tacticals.ml b/tactics/tacticals.ml index 9cf3c4187..94f22f903 100644 --- a/tactics/tacticals.ml +++ b/tactics/tacticals.ml @@ -158,7 +158,7 @@ type branch_args = { type branch_assumptions = { ba : branch_args; (* the branch args *) - assums : Context.Named.t} (* the list of assumptions introduced *) + assums : named_context} (* the list of assumptions introduced *) open Misctypes @@ -625,7 +625,6 @@ module New = struct (* c should be of type A1->.. An->B with B an inductive definition *) let general_elim_then_using mk_elim isrec allnames tac predicate ind (c, t) = - let open EConstr in Proofview.Goal.nf_enter { enter = begin fun gl -> let sigma, elim = Tacmach.New.of_old (mk_elim ind) gl in Proofview.tclTHEN (Proofview.Unsafe.tclEVARS sigma) diff --git a/tactics/tacticals.mli b/tactics/tacticals.mli index 2b07d937e..e9f623100 100644 --- a/tactics/tacticals.mli +++ b/tactics/tacticals.mli @@ -60,29 +60,29 @@ val tclIFTHENTRYELSEMUST : tactic -> tactic -> tactic val onNthHypId : int -> (Id.t -> tactic) -> tactic val onNthHyp : int -> (constr -> tactic) -> tactic -val onNthDecl : int -> (Context.Named.Declaration.t -> tactic) -> tactic +val onNthDecl : int -> (named_declaration -> tactic) -> tactic val onLastHypId : (Id.t -> tactic) -> tactic val onLastHyp : (constr -> tactic) -> tactic -val onLastDecl : (Context.Named.Declaration.t -> tactic) -> tactic +val onLastDecl : (named_declaration -> tactic) -> tactic val onNLastHypsId : int -> (Id.t list -> tactic) -> tactic val onNLastHyps : int -> (constr list -> tactic) -> tactic -val onNLastDecls : int -> (Context.Named.t -> tactic) -> tactic +val onNLastDecls : int -> (named_context -> tactic) -> tactic val lastHypId : goal sigma -> Id.t val lastHyp : goal sigma -> constr -val lastDecl : goal sigma -> Context.Named.Declaration.t +val lastDecl : goal sigma -> named_declaration val nLastHypsId : int -> goal sigma -> Id.t list val nLastHyps : int -> goal sigma -> constr list -val nLastDecls : int -> goal sigma -> Context.Named.t +val nLastDecls : int -> goal sigma -> named_context -val afterHyp : Id.t -> goal sigma -> Context.Named.t +val afterHyp : Id.t -> goal sigma -> named_context val ifOnHyp : (Id.t * types -> bool) -> (Id.t -> tactic) -> (Id.t -> tactic) -> Id.t -> tactic -val onHyps : (goal sigma -> Context.Named.t) -> - (Context.Named.t -> tactic) -> tactic +val onHyps : (goal sigma -> named_context) -> + (named_context -> tactic) -> tactic (** {6 Tacticals applying to goal components } *) @@ -110,7 +110,7 @@ type branch_args = private { type branch_assumptions = private { ba : branch_args; (** the branch args *) - assums : Context.Named.t} (** the list of assumptions introduced *) + assums : named_context} (** the list of assumptions introduced *) (** [get_and_check_or_and_pattern loc pats branchsign] returns an appropriate error message if |pats| <> |branchsign|; extends them if no pattern is given @@ -229,7 +229,7 @@ module New : sig val tclTIMEOUT : int -> unit tactic -> unit tactic val tclTIME : string option -> 'a tactic -> 'a tactic - val nLastDecls : ([ `NF ], 'r) Proofview.Goal.t -> int -> Context.Named.t + val nLastDecls : ([ `NF ], 'r) Proofview.Goal.t -> int -> named_context val ifOnHyp : (identifier * types -> bool) -> (identifier -> unit Proofview.tactic) -> (identifier -> unit Proofview.tactic) -> @@ -238,11 +238,11 @@ module New : sig val onNthHypId : int -> (identifier -> unit tactic) -> unit tactic val onLastHypId : (identifier -> unit tactic) -> unit tactic val onLastHyp : (constr -> unit tactic) -> unit tactic - val onLastDecl : (Context.Named.Declaration.t -> unit tactic) -> unit tactic + val onLastDecl : (named_declaration -> unit tactic) -> unit tactic - val onHyps : ([ `NF ], Context.Named.t) Proofview.Goal.enter -> - (Context.Named.t -> unit tactic) -> unit tactic - val afterHyp : Id.t -> (Context.Named.t -> unit tactic) -> unit tactic + val onHyps : ([ `NF ], named_context) Proofview.Goal.enter -> + (named_context -> unit tactic) -> unit tactic + val afterHyp : Id.t -> (named_context -> unit tactic) -> unit tactic val tryAllHyps : (identifier -> unit tactic) -> unit tactic val tryAllHypsAndConcl : (identifier option -> unit tactic) -> unit tactic diff --git a/tactics/tactics.ml b/tactics/tactics.ml index 8260c14ad..4bf848a9c 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -15,6 +15,7 @@ open Names open Nameops open Term open Termops +open Environ open EConstr open Vars open Find_subterm @@ -22,7 +23,6 @@ open Namegen open Declarations open Inductiveops open Reductionops -open Environ open Globnames open Evd open Pfedit @@ -170,26 +170,6 @@ let _ = (* Primitive tactics *) (******************************************) -let local_assum (na, t) = - let open Context.Rel.Declaration in - let inj = EConstr.Unsafe.to_constr in - LocalAssum (na, inj t) - -let local_def (na, b, t) = - let open Context.Rel.Declaration in - let inj = EConstr.Unsafe.to_constr in - LocalDef (na, inj b, inj t) - -let nlocal_assum (na, t) = - let open Context.Named.Declaration in - let inj = EConstr.Unsafe.to_constr in - LocalAssum (na, inj t) - -let nlocal_def (na, b, t) = - let open Context.Named.Declaration in - let inj = EConstr.Unsafe.to_constr in - LocalDef (na, inj b, inj t) - (** This tactic creates a partial proof realizing the introduction rule, but does not check anything. *) let unsafe_intro env store decl b = @@ -217,8 +197,8 @@ let introduction ?(check=true) id = in let open Context.Named.Declaration in match EConstr.kind sigma concl with - | Prod (_, t, b) -> unsafe_intro env store (nlocal_assum (id, t)) b - | LetIn (_, c, t, b) -> unsafe_intro env store (nlocal_def (id, c, t)) b + | Prod (_, t, b) -> unsafe_intro env store (LocalAssum (id, t)) b + | LetIn (_, c, t, b) -> unsafe_intro env store (LocalDef (id, c, t)) b | _ -> raise (RefinerError IntroNeedsProduct) end } @@ -321,7 +301,6 @@ let clear_gen fail = function try clear_hyps_in_evi env evdref (named_context_val env) concl ids with Evarutil.ClearDependencyError (id,err) -> fail env sigma id err in - let concl = EConstr.of_constr concl in let env = reset_with_named_context hyps env in let tac = Refine.refine ~unsafe:true { run = fun sigma -> Evarutil.new_evar env sigma ~principal:true concl @@ -397,18 +376,16 @@ let rename_hyp repl = with Not_found -> () in (** All is well *) - let make_subst (src, dst) = (src, Constr.mkVar dst) in + let make_subst (src, dst) = (src, mkVar dst) in let subst = List.map make_subst repl in - let subst c = CVars.replace_vars subst c in + let subst c = Vars.replace_vars subst c in let map decl = decl |> NamedDecl.map_id (fun id -> try List.assoc_f Id.equal id repl with Not_found -> id) |> NamedDecl.map_constr subst in let nhyps = List.map map hyps in - let concl = EConstr.Unsafe.to_constr concl in let nconcl = subst concl in - let nconcl = EConstr.of_constr nconcl in - let nctx = Environ.val_of_named_context nhyps in + let nctx = val_of_named_context nhyps in let instance = List.map (NamedDecl.get_id %> mkVar) hyps in Refine.refine ~unsafe:true { run = begin fun sigma -> Evarutil.new_evar_instance nctx sigma nconcl ~principal:true ~store instance @@ -435,11 +412,14 @@ let id_of_name_with_default id = function let default_id_of_sort s = if Sorts.is_small s then default_small_ident else default_type_ident +let id_of_name_using_hdchar env c name = + id_of_name_using_hdchar env (EConstr.Unsafe.to_constr c) name + let default_id env sigma decl = let open Context.Rel.Declaration in match decl with | LocalAssum (name,t) -> - let dft = default_id_of_sort (Retyping.get_sort_of env sigma (EConstr.of_constr t)) in + let dft = default_id_of_sort (Retyping.get_sort_of env sigma t) in id_of_name_with_default dft name | LocalDef (name,b,_) -> id_of_name_using_hdchar env b name @@ -478,7 +458,7 @@ let find_name mayrepl decl naming gl = match naming with let assert_before_then_gen b naming t tac = let open Context.Rel.Declaration in Proofview.Goal.enter { enter = begin fun gl -> - let id = find_name b (local_assum (Anonymous,t)) naming gl in + let id = find_name b (LocalAssum (Anonymous,t)) naming gl in Tacticals.New.tclTHENLAST (Proofview.V82.tactic (fun gl -> @@ -497,7 +477,7 @@ let assert_before_replacing id = assert_before_gen true (NamingMustBe (dloc,id)) let assert_after_then_gen b naming t tac = let open Context.Rel.Declaration in Proofview.Goal.enter { enter = begin fun gl -> - let id = find_name b (local_assum (Anonymous,t)) naming gl in + let id = find_name b (LocalAssum (Anonymous,t)) naming gl in Tacticals.New.tclTHENFIRST (Proofview.V82.tactic (fun gl -> @@ -534,7 +514,7 @@ let rec check_mutind env sigma k cl = match EConstr.kind sigma (strip_outer_cast with Not_found -> error "Cannot do a fixpoint on a non inductive type." else let open Context.Rel.Declaration in - check_mutind (push_rel (local_assum (na, c1)) env) sigma (pred k) b + check_mutind (push_rel (LocalAssum (na, c1)) env) sigma (pred k) b | _ -> error "Not enough products." (* Refine as a fixpoint *) @@ -548,13 +528,14 @@ let mutual_fix f n rest j = Proofview.Goal.nf_enter { enter = begin fun gl -> let rec mk_sign sign = function | [] -> sign | (f, n, ar) :: oth -> + let open Context.Named.Declaration in let (sp', u') = check_mutind env sigma n ar in if not (eq_mind sp sp') then error "Fixpoints should be on the same mutual inductive declaration."; if mem_named_context_val f sign then user_err ~hdr:"Logic.prim_refiner" (str "Name " ++ pr_id f ++ str " already used in the environment"); - mk_sign (push_named_context_val (nlocal_assum (f, ar)) sign) oth + mk_sign (push_named_context_val (LocalAssum (f, ar)) sign) oth in let nenv = reset_with_named_context (mk_sign (named_context_val env) all) env in Refine.refine { run = begin fun sigma -> @@ -584,7 +565,8 @@ let rec check_is_mutcoind env sigma cl = let b = whd_all env sigma cl in match EConstr.kind sigma b with | Prod (na, c1, b) -> - check_is_mutcoind (push_rel (local_assum (na,c1)) env) sigma b + let open Context.Rel.Declaration in + check_is_mutcoind (push_rel (LocalAssum (na,c1)) env) sigma b | _ -> try let _ = find_coinductive env sigma b in () @@ -602,9 +584,10 @@ let mutual_cofix f others j = Proofview.Goal.nf_enter { enter = begin fun gl -> let rec mk_sign sign = function | [] -> sign | (f, ar) :: oth -> + let open Context.Named.Declaration in if mem_named_context_val f sign then error "Name already used in the environment."; - mk_sign (push_named_context_val (nlocal_assum (f, ar)) sign) oth + mk_sign (push_named_context_val (LocalAssum (f, ar)) sign) oth in let nenv = reset_with_named_context (mk_sign (named_context_val env) all) env in Refine.refine { run = begin fun sigma -> @@ -640,16 +623,13 @@ let pf_reduce_decl redfun where decl gl = let redfun' c = Tacmach.New.pf_apply redfun gl c in match decl with | LocalAssum (id,ty) -> - let ty = EConstr.of_constr ty in if where == InHypValueOnly then user_err (pr_id id ++ str " has no value."); - nlocal_assum (id,redfun' ty) + LocalAssum (id,redfun' ty) | LocalDef (id,b,ty) -> - let b = EConstr.of_constr b in - let ty = EConstr.of_constr ty in let b' = if where != InHypTypeOnly then redfun' b else b in let ty' = if where != InHypValueOnly then redfun' ty else ty in - nlocal_def (id,b',ty') + LocalDef (id,b',ty') (* Possibly equip a reduction with the occurrences mentioned in an occurrence clause *) @@ -744,17 +724,14 @@ let pf_e_reduce_decl redfun where decl gl = let redfun sigma c = redfun.e_redfun (Tacmach.New.pf_env gl) sigma c in match decl with | LocalAssum (id,ty) -> - let ty = EConstr.of_constr ty in if where == InHypValueOnly then user_err (pr_id id ++ str " has no value."); let Sigma (ty', sigma, p) = redfun sigma ty in - Sigma (nlocal_assum (id, ty'), sigma, p) + Sigma (LocalAssum (id, ty'), sigma, p) | LocalDef (id,b,ty) -> - let b = EConstr.of_constr b in - let ty = EConstr.of_constr ty in let Sigma (b', sigma, p) = if where != InHypTypeOnly then redfun sigma b else Sigma.here b sigma in let Sigma (ty', sigma, q) = if where != InHypValueOnly then redfun sigma ty else Sigma.here ty sigma in - Sigma (nlocal_def (id, b', ty'), sigma, p +> q) + Sigma (LocalDef (id, b', ty'), sigma, p +> q) let e_reduct_in_concl ~check (redfun, sty) = Proofview.Goal.nf_s_enter { s_enter = begin fun gl -> @@ -787,21 +764,18 @@ let e_pf_change_decl (redfun : bool -> e_reduction_function) where decl env sigm let open Context.Named.Declaration in match decl with | LocalAssum (id,ty) -> - let ty = EConstr.of_constr ty in if where == InHypValueOnly then user_err (pr_id id ++ str " has no value."); let Sigma (ty', sigma, p) = (redfun false).e_redfun env sigma ty in - Sigma (nlocal_assum (id, ty'), sigma, p) + Sigma (LocalAssum (id, ty'), sigma, p) | LocalDef (id,b,ty) -> - let b = EConstr.of_constr b in - let ty = EConstr.of_constr ty in let Sigma (b', sigma, p) = if where != InHypTypeOnly then (redfun true).e_redfun env sigma b else Sigma.here b sigma in let Sigma (ty', sigma, q) = if where != InHypValueOnly then (redfun false).e_redfun env sigma ty else Sigma.here ty sigma in - Sigma (nlocal_def (id,b',ty'), sigma, p +> q) + Sigma (LocalDef (id,b',ty'), sigma, p +> q) let e_change_in_hyp redfun (id,where) = Proofview.Goal.s_enter { s_enter = begin fun gl -> @@ -974,10 +948,10 @@ let rec intro_then_gen name_flag move_flag force_flag dep_flag tac = let concl = Proofview.Goal.concl (Proofview.Goal.assume gl) in match EConstr.kind sigma concl with | Prod (name,t,u) when not dep_flag || not (noccurn sigma 1 u) -> - let name = find_name false (local_assum (name,t)) name_flag gl in + let name = find_name false (LocalAssum (name,t)) name_flag gl in build_intro_tac name move_flag tac | LetIn (name,b,t,u) when not dep_flag || not (noccurn sigma 1 u) -> - let name = find_name false (local_def (name,b,t)) name_flag gl in + let name = find_name false (LocalDef (name,b,t)) name_flag gl in build_intro_tac name move_flag tac | _ -> begin if not force_flag then Proofview.tclZERO (RefinerError IntroNeedsProduct) @@ -1382,11 +1356,11 @@ let enforce_prop_bound_names rename tac = Name (add_suffix Namegen.default_prop_ident s) else na in - mkProd (na,t,aux (push_rel (local_assum (na,t)) env) sigma (i-1) t') + mkProd (na,t,aux (push_rel (LocalAssum (na,t)) env) sigma (i-1) t') | Prod (Anonymous,t,t') -> - mkProd (Anonymous,t,aux (push_rel (local_assum (Anonymous,t)) env) sigma (i-1) t') + mkProd (Anonymous,t,aux (push_rel (LocalAssum (Anonymous,t)) env) sigma (i-1) t') | LetIn (na,c,t,t') -> - mkLetIn (na,c,t,aux (push_rel (local_def (na,c,t)) env) sigma (i-1) t') + mkLetIn (na,c,t,aux (push_rel (LocalDef (na,c,t)) env) sigma (i-1) t') | _ -> assert false in let rename_branch i = Proofview.Goal.nf_enter { enter = begin fun gl -> @@ -1619,11 +1593,11 @@ let make_projection env sigma params cstr sign elim i n c u = let elim = match elim with | NotADefinedRecordUseScheme elim -> (* bugs: goes from right to left when i increases! *) - let decl = List.nth cstr.cs_args i in + let cs_args = List.map (fun d -> map_rel_decl EConstr.of_constr d) cstr.cs_args in + let decl = List.nth cs_args i in let t = RelDecl.get_type decl in - let t = EConstr.of_constr t in - let b = match decl with LocalAssum _ -> mkRel (i+1) | LocalDef (_,b,_) -> EConstr.of_constr b in - let branch = it_mkLambda_or_LetIn b cstr.cs_args in + let b = match decl with LocalAssum _ -> mkRel (i+1) | LocalDef (_,b,_) -> b in + let branch = it_mkLambda_or_LetIn b cs_args in if (* excludes dependent projection types *) noccur_between sigma 1 (n-i-1) t @@ -1890,7 +1864,7 @@ let apply_in_once sidecond_first with_delta with_destruct with_evars naming if with_delta then default_unify_flags () else default_no_delta_unify_flags () in let t' = Tacmach.New.pf_get_hyp_typ id gl in let innerclause = mk_clenv_from_env env sigma (Some 0) (mkVar id,t') in - let targetid = find_name true (local_assum (Anonymous,t')) naming gl in + let targetid = find_name true (LocalAssum (Anonymous,t')) naming gl in let rec aux idstoclear with_destruct c = Proofview.Goal.enter { enter = begin fun gl -> let env = Proofview.Goal.env gl in @@ -2017,7 +1991,6 @@ let assumption = else Tacticals.New.tclZEROMSG (str "No such assumption.") | decl::rest -> let t = NamedDecl.get_type decl in - let t = EConstr.of_constr t in let concl = Proofview.Goal.concl gl in let sigma = Tacmach.New.project gl in let (sigma, is_same_type) = @@ -2058,13 +2031,13 @@ let check_is_type env sigma ty = let check_decl env sigma decl = let open Context.Named.Declaration in - let ty = EConstr.of_constr (NamedDecl.get_type decl) in + let ty = NamedDecl.get_type decl in let evdref = ref sigma in try let _ = Typing.e_sort_of env evdref ty in let _ = match decl with | LocalAssum _ -> () - | LocalDef (_,c,_) -> Typing.e_check env evdref (EConstr.of_constr c) ty + | LocalDef (_,c,_) -> Typing.e_check env evdref c ty in !evdref with e when CErrors.noncritical e -> @@ -2146,6 +2119,7 @@ let keep hyps = let sigma = Tacmach.New.project gl in let cl,_ = fold_named_context_reverse (fun (clear,keep) decl -> + let decl = map_named_decl EConstr.of_constr decl in let hyp = NamedDecl.get_id decl in if Id.List.mem hyp hyps || List.exists (occur_var_in_decl env sigma hyp) keep @@ -2692,6 +2666,7 @@ let insert_before decls lasthyp env = | Some id -> Environ.fold_named_context (fun _ d env -> + let d = map_named_decl EConstr.of_constr d in let env = if Id.equal id (NamedDecl.get_id d) then push_named_context decls env else env in push_named d env) ~init:(reset_context env) env @@ -2701,8 +2676,8 @@ let insert_before decls lasthyp env = let mkletin_goal env sigma store with_eq dep (id,lastlhyp,ccl,c) ty = let open Context.Named.Declaration in let t = match ty with Some t -> t | _ -> typ_of env sigma c in - let decl = if dep then nlocal_def (id,c,t) - else nlocal_assum (id,t) + let decl = if dep then LocalDef (id,c,t) + else LocalAssum (id,t) in match with_eq with | Some (lr,(loc,ido)) -> @@ -2721,7 +2696,7 @@ let mkletin_goal env sigma store with_eq dep (id,lastlhyp,ccl,c) ty = let refl = EConstr.of_constr refl in let eq = applist (eq,args) in let refl = applist (refl, [t;mkVar id]) in - let newenv = insert_before [nlocal_assum (heq,eq); decl] lastlhyp env in + let newenv = insert_before [LocalAssum (heq,eq); decl] lastlhyp env in let Sigma (x, sigma, r) = new_evar newenv sigma ~principal:true ~store ccl in Sigma (mkNamedLetIn id c t (mkNamedLetIn heq refl eq x), sigma, p +> q +> r) | None -> @@ -2822,8 +2797,8 @@ let generalize_goal_gen env sigma ids i ((occs,c,b),na) t cl = let cl',sigma' = subst_closed_term_occ env sigma (AtOccs occs) c (it_mkProd_or_LetIn cl newdecls) in let na = generalized_name sigma c t ids cl' na in let decl = match b with - | None -> local_assum (na,t) - | Some b -> local_def (na,b,t) + | None -> LocalAssum (na,t) + | Some b -> LocalDef (na,b,t) in mkProd_or_LetIn decl cl', sigma' @@ -2838,7 +2813,7 @@ let old_generalize_dep ?(with_let=false) c gl = let sign = pf_hyps gl in let sigma = project gl in let init_ids = ids_of_named_context (Global.named_context()) in - let seek (d:Context.Named.Declaration.t) (toquant:Context.Named.t) = + let seek (d:named_declaration) (toquant:named_context) = if List.exists (fun d' -> occur_var_in_decl env sigma (NamedDecl.get_id d') d) toquant || dependent_in_decl sigma c d then d::toquant @@ -2862,7 +2837,6 @@ let old_generalize_dep ?(with_let=false) c gl = | _ -> None else None in - let body = Option.map EConstr.of_constr body in let cl'',evd = generalize_goal gl 0 ((AllOccurrences,c,body),Anonymous) (cl',project gl) in (** Check that the generalization is indeed well-typed *) @@ -3256,7 +3230,7 @@ let atomize_param_of_ind_then (indref,nparams,_) hyp0 tac = | Var id -> id | _ -> let type_of = Tacmach.New.pf_unsafe_type_of gl in - id_of_name_using_hdchar (Global.env()) (EConstr.Unsafe.to_constr (type_of c)) Anonymous in + id_of_name_using_hdchar (Global.env()) (type_of c) Anonymous in let x = fresh_id_in_env avoid id env in Tacticals.New.tclTHEN (letin_tac None (Name x) c None allHypsAndConcl) @@ -3342,6 +3316,7 @@ let cook_sign hyp0_opt inhyps indvars env sigma = let before = ref true in let maindep = ref false in let seek_deps env decl rhyp = + let decl = map_named_decl EConstr.of_constr decl in let hyp = NamedDecl.get_id decl in if (match hyp0_opt with Some hyp0 -> Id.equal hyp hyp0 | _ -> false) then begin @@ -3434,15 +3409,15 @@ type elim_scheme = { elimc: constr with_bindings option; elimt: types; indref: global_reference option; - params: Context.Rel.t; (* (prm1,tprm1);(prm2,tprm2)...(prmp,tprmp) *) + params: rel_context; (* (prm1,tprm1);(prm2,tprm2)...(prmp,tprmp) *) nparams: int; (* number of parameters *) - predicates: Context.Rel.t; (* (Qq, (Tq_1 -> Tq_2 ->...-> Tq_nq)), (Q1,...) *) + predicates: rel_context; (* (Qq, (Tq_1 -> Tq_2 ->...-> Tq_nq)), (Q1,...) *) npredicates: int; (* Number of predicates *) - branches: Context.Rel.t; (* branchr,...,branch1 *) + branches: rel_context; (* branchr,...,branch1 *) nbranches: int; (* Number of branches *) - args: Context.Rel.t; (* (xni, Ti_ni) ... (x1, Ti_1) *) + args: rel_context; (* (xni, Ti_ni) ... (x1, Ti_1) *) nargs: int; (* number of arguments *) - indarg: Context.Rel.Declaration.t option; (* Some (H,I prm1..prmp x1...xni) + indarg: rel_declaration option; (* Some (H,I prm1..prmp x1...xni) if HI is in premisses, None otherwise *) concl: types; (* Qi x1...xni HI (f...), HI and (f...) are optional and mutually exclusive *) @@ -3586,10 +3561,10 @@ let make_abstract_generalize env id typ concl dep ctx body c eqs args refls = in (* Abstract by equalities *) let eqs = lift_togethern 1 eqs in (* lift together and past genarg *) - let abseqs = it_mkProd_or_LetIn (lift eqslen abshypeq) (List.map (fun x -> local_assum (Anonymous, x)) eqs) in + let abseqs = it_mkProd_or_LetIn (lift eqslen abshypeq) (List.map (fun x -> LocalAssum (Anonymous, x)) eqs) in let decl = match body with - | None -> local_assum (Name id, c) - | Some body -> local_def (Name id, body, c) + | None -> LocalAssum (Name id, c) + | Some body -> LocalDef (Name id, body, c) in (* Abstract by the "generalized" hypothesis. *) let genarg = mkProd_or_LetIn decl abseqs in @@ -3668,7 +3643,6 @@ let abstract_args gl generalize_vars dep id defined f args = let decl = List.hd rel in RelDecl.get_name decl, RelDecl.get_type decl, c in - let ty = EConstr.of_constr ty in let argty = Tacmach.pf_unsafe_type_of gl arg in let sigma', ty = Evarsolve.refresh_universes (Some true) env !sigma ty in let () = sigma := sigma' in @@ -3681,7 +3655,7 @@ let abstract_args gl generalize_vars dep id defined f args = Id.Set.add id nongenvars, Id.Set.remove id vars, env) | _ -> let name = get_id name in - let decl = local_assum (Name name, ty) in + let decl = LocalAssum (Name name, ty) in let ctx = decl :: ctx in let c' = mkApp (lift 1 c, [|mkRel 1|]) in let args = arg :: args in @@ -3739,10 +3713,9 @@ let abstract_generalize ?(generalize_vars=true) ?(force_dep=false) id = let (f, args, def, id, oldid) = let oldid = Tacmach.New.pf_get_new_id id gl in match Tacmach.New.pf_get_hyp id gl with - | LocalAssum (_,t) -> let f, args = decompose_app sigma (EConstr.of_constr t) in + | LocalAssum (_,t) -> let f, args = decompose_app sigma t in (f, args, false, id, oldid) | LocalDef (_,t,_) -> - let t = EConstr.of_constr t in let f, args = decompose_app sigma t in (f, args, true, id, oldid) in @@ -3809,13 +3782,13 @@ let specialize_eqs id gl = if in_eqs then acc, in_eqs, ctx, ty else let e = e_new_evar (push_rel_context ctx env) evars t in - aux false (local_def (na,e,t) :: ctx) (mkApp (lift 1 acc, [| mkRel 1 |])) b) + aux false (LocalDef (na,e,t) :: ctx) (mkApp (lift 1 acc, [| mkRel 1 |])) b) | t -> acc, in_eqs, ctx, ty in let acc, worked, ctx, ty = aux false [] (mkVar id) ty in let ctx' = nf_rel_context_evar !evars ctx in let ctx'' = List.map (function - | LocalDef (n,k,t) when isEvar !evars (EConstr.of_constr k) -> LocalAssum (n,t) + | LocalDef (n,k,t) when isEvar !evars k -> LocalAssum (n,t) | decl -> decl) ctx' in let ty' = it_mkProd_or_LetIn ty ctx'' in @@ -3855,13 +3828,13 @@ let decompose_paramspred_branch_args sigma elimt = | Prod(nme,tpe,elimt') -> let hd_tpe,_ = decompose_app sigma (snd (decompose_prod_assum sigma tpe)) in if not (occur_rel sigma 1 elimt') && isRel sigma hd_tpe - then cut_noccur elimt' (local_assum (nme,tpe)::acc2) + then cut_noccur elimt' (LocalAssum (nme,tpe)::acc2) else let acc3,ccl = decompose_prod_assum sigma elimt in acc2 , acc3 , ccl | App(_, _) | Rel _ -> acc2 , [] , elimt | _ -> error_ind_scheme "" in let rec cut_occur elimt acc1 = match EConstr.kind sigma elimt with - | Prod(nme,tpe,c) when occur_rel sigma 1 c -> cut_occur c (local_assum (nme,tpe)::acc1) + | Prod(nme,tpe,c) when occur_rel sigma 1 c -> cut_occur c (LocalAssum (nme,tpe)::acc1) | Prod(nme,tpe,c) -> let acc2,acc3,ccl = cut_noccur elimt [] in acc1,acc2,acc3,ccl | App(_, _) | Rel _ -> acc1,[],[],elimt | _ -> error_ind_scheme "" in @@ -3939,7 +3912,6 @@ let compute_elim_sig sigma ?elimc elimt = match List.hd args_indargs with | LocalDef (hiname,_,hi) -> error_ind_scheme "" | LocalAssum (hiname,hi) -> - let hi = EConstr.of_constr hi in let hi_ind, hi_args = decompose_app sigma hi in let hi_is_ind = (* hi est d'un type globalisable *) match EConstr.kind sigma hi_ind with @@ -3965,7 +3937,6 @@ let compute_elim_sig sigma ?elimc elimt = | None -> !res (* No indref *) | Some (LocalDef _) -> error_ind_scheme "" | Some (LocalAssum (_,ind)) -> - let ind = EConstr.of_constr ind in let indhd,indargs = decompose_app sigma ind in try {!res with indref = Some (fst (Termops.global_of_constr sigma indhd)) } with e when CErrors.noncritical e -> @@ -3983,7 +3954,6 @@ let compute_scheme_signature evd scheme names_info ind_type_guess = let cond hd = EConstr.eq_constr evd hd ind_type_guess && not scheme.farg_in_concl in (cond, fun _ _ -> ()) | Some (LocalAssum (_,ind)) -> (* Standard scheme from an inductive type *) - let ind = EConstr.of_constr ind in let indhd,indargs = decompose_app evd ind in let cond hd = EConstr.eq_constr evd hd indhd in let check_concl is_pred p = @@ -4016,7 +3986,6 @@ let compute_scheme_signature evd scheme names_info ind_type_guess = let rec find_branches p lbrch = match lbrch with | LocalAssum (_,t) :: brs -> - let t = EConstr.of_constr t in (try let lchck_brch = check_branch p t in let n = List.fold_left @@ -4123,7 +4092,7 @@ let get_eliminator elim dep s gl = | ElimOver (isrec,id) -> let evd, (elimc,elimt),_ as elims = guess_elim isrec dep s id gl in let _, (l, s) = compute_elim_signature elims id in - let branchlengthes = List.map (fun d -> assert (RelDecl.is_local_assum d); pi1 (decompose_prod_letin (Tacmach.New.project gl) (EConstr.of_constr (RelDecl.get_type d)))) + let branchlengthes = List.map (fun d -> assert (RelDecl.is_local_assum d); pi1 (decompose_prod_letin (Tacmach.New.project gl) (RelDecl.get_type d))) (List.rev s.branches) in evd, isrec, ({elimindex = None; elimbody = elimc; elimrename = Some (isrec,Array.of_list branchlengthes)}, elimt), l @@ -4465,7 +4434,6 @@ let induction_gen clear_flag isrec with_evars elim declaring the induction argument as a new local variable *) let id = (* Type not the right one if partially applied but anyway for internal use*) - let t = EConstr.Unsafe.to_constr t in let x = id_of_name_using_hdchar (Global.env()) t Anonymous in new_fresh_id [] x gl in let info_arg = (is_arg_pure_hyp, not enough_applied) in @@ -4503,7 +4471,7 @@ let induction_gen_l isrec with_evars elim names lc = let type_of = Tacmach.New.pf_unsafe_type_of gl in let sigma = Tacmach.New.project gl in let x = - id_of_name_using_hdchar (Global.env()) (EConstr.Unsafe.to_constr (type_of c)) Anonymous in + id_of_name_using_hdchar (Global.env()) (type_of c) Anonymous in let id = new_fresh_id [] x gl in let newl' = List.map (fun r -> replace_term sigma c (mkVar id) r) l' in @@ -4863,6 +4831,9 @@ let intros_transitivity n = Tacticals.New.tclTHEN intros (transitivity_gen n) (** d1 is the section variable in the global context, d2 in the goal context *) let interpretable_as_section_decl evd d1 d2 = let open Context.Named.Declaration in + let e_eq_constr_univs sigma c1 c2 = + e_eq_constr_univs sigma (EConstr.Unsafe.to_constr c1) (EConstr.Unsafe.to_constr c2) + in match d2, d1 with | LocalDef _, LocalAssum _ -> false | LocalDef (_,b1,t1), LocalDef (_,b2,t2) -> diff --git a/tactics/tactics.mli b/tactics/tactics.mli index b0d9dcb1c..0087d607d 100644 --- a/tactics/tactics.mli +++ b/tactics/tactics.mli @@ -35,9 +35,9 @@ val is_quantified_hypothesis : Id.t -> ([`NF],'b) Proofview.Goal.t -> bool val introduction : ?check:bool -> Id.t -> unit Proofview.tactic val convert_concl : ?check:bool -> types -> cast_kind -> unit Proofview.tactic -val convert_hyp : ?check:bool -> Context.Named.Declaration.t -> unit Proofview.tactic +val convert_hyp : ?check:bool -> named_declaration -> unit Proofview.tactic val convert_concl_no_check : types -> cast_kind -> unit Proofview.tactic -val convert_hyp_no_check : Context.Named.Declaration.t -> unit Proofview.tactic +val convert_hyp_no_check : named_declaration -> unit Proofview.tactic val mutual_fix : Id.t -> int -> (Id.t * int * constr) list -> int -> unit Proofview.tactic val fix : Id.t option -> int -> unit Proofview.tactic @@ -51,7 +51,7 @@ val convert_leq : constr -> constr -> unit Proofview.tactic val fresh_id_in_env : Id.t list -> Id.t -> env -> Id.t val fresh_id : Id.t list -> Id.t -> goal sigma -> Id.t -val find_intro_names : Context.Rel.t -> goal sigma -> Id.t list +val find_intro_names : rel_context -> goal sigma -> Id.t list val intro : unit Proofview.tactic val introf : unit Proofview.tactic @@ -185,7 +185,7 @@ val revert : Id.t list -> unit Proofview.tactic (** {6 Resolution tactics. } *) val apply_type : constr -> constr list -> unit Proofview.tactic -val bring_hyps : Context.Named.t -> unit Proofview.tactic +val bring_hyps : named_context -> unit Proofview.tactic val apply : constr -> unit Proofview.tactic val eapply : constr -> unit Proofview.tactic @@ -244,15 +244,15 @@ type elim_scheme = { elimc: constr with_bindings option; elimt: types; indref: global_reference option; - params: Context.Rel.t; (** (prm1,tprm1);(prm2,tprm2)...(prmp,tprmp) *) + params: rel_context; (** (prm1,tprm1);(prm2,tprm2)...(prmp,tprmp) *) nparams: int; (** number of parameters *) - predicates: Context.Rel.t; (** (Qq, (Tq_1 -> Tq_2 ->...-> Tq_nq)), (Q1,...) *) + predicates: rel_context; (** (Qq, (Tq_1 -> Tq_2 ->...-> Tq_nq)), (Q1,...) *) npredicates: int; (** Number of predicates *) - branches: Context.Rel.t; (** branchr,...,branch1 *) + branches: rel_context; (** branchr,...,branch1 *) nbranches: int; (** Number of branches *) - args: Context.Rel.t; (** (xni, Ti_ni) ... (x1, Ti_1) *) + args: rel_context; (** (xni, Ti_ni) ... (x1, Ti_1) *) nargs: int; (** number of arguments *) - indarg: Context.Rel.Declaration.t option; (** Some (H,I prm1..prmp x1...xni) + indarg: rel_declaration option; (** Some (H,I prm1..prmp x1...xni) if HI is in premisses, None otherwise *) concl: types; (** Qi x1...xni HI (f...), HI and (f...) are optional and mutually exclusive *) diff --git a/tactics/term_dnet.ml b/tactics/term_dnet.ml index 219abb7fd..2c863c42a 100644 --- a/tactics/term_dnet.ml +++ b/tactics/term_dnet.ml @@ -344,7 +344,7 @@ struct ) (pr_dconstr pr_term_pattern) p*) let search_pat cpat dpat dn = - let whole_c = cpat in + let whole_c = EConstr.of_constr cpat in (* if we are at the root, add an empty context *) let dpat = under_prod (empty_ctx dpat) in TDnet.Idset.fold @@ -352,9 +352,8 @@ struct let c_id = Opt.reduce (Ident.constr_of id) in let c_id = EConstr.of_constr c_id in let (ctx,wc) = - try Termops.align_prod_letin Evd.empty (EConstr.of_constr whole_c) c_id (** FIXME *) + try Termops.align_prod_letin Evd.empty whole_c c_id (** FIXME *) with Invalid_argument _ -> [],c_id in - let wc = EConstr.Unsafe.to_constr wc in let wc,whole_c = if Opt.direction then whole_c,wc else wc,whole_c in try let _ = Termops.filtering Evd.empty ctx Reduction.CUMUL wc whole_c in diff --git a/toplevel/classes.ml b/toplevel/classes.ml index 7a5d0ed81..165663ef6 100644 --- a/toplevel/classes.ml +++ b/toplevel/classes.ml @@ -150,6 +150,7 @@ let new_instance ?(abstract=false) ?(global=false) ?(refine= !refine_instance) p in let k, u, cty, ctx', ctx, len, imps, subst = let impls, ((env', ctx), imps) = interp_context_evars env evars ctx in + let ctx = List.map (fun d -> Termops.map_rel_decl EConstr.Unsafe.to_constr d) ctx in let c', imps' = interp_type_evars_impls ~impls env' evars tclass in let c' = EConstr.Unsafe.to_constr c' in let len = List.length ctx in @@ -363,6 +364,7 @@ let context poly l = let env = Global.env() in let evars = ref (Evd.from_env env) in let _, ((env', fullctx), impls) = interp_context_evars env evars l in + let fullctx = List.map (fun d -> Termops.map_rel_decl EConstr.Unsafe.to_constr d) fullctx in let subst = Evarutil.evd_comb0 Evarutil.nf_evars_and_universes evars in let fullctx = Context.Rel.map subst fullctx in let ce t = Pretyping.check_evars env Evd.empty !evars (EConstr.of_constr t) in diff --git a/toplevel/command.ml b/toplevel/command.ml index 44fc4eb1a..b1415515f 100644 --- a/toplevel/command.ml +++ b/toplevel/command.ml @@ -95,6 +95,7 @@ let interp_definition pl bl p red_option c ctypopt = let ctx = Evd.make_evar_universe_context env pl in let evdref = ref (Evd.from_ctx ctx) in let impls, ((env_bl, ctx), imps1) = interp_context_evars env evdref bl in + let ctx = List.map (fun d -> map_rel_decl EConstr.Unsafe.to_constr d) ctx in let nb_args = List.length ctx in let imps,pl,ce = match ctypopt with @@ -581,6 +582,7 @@ let interp_mutual_inductive (paramsl,indl) notations poly prv finite = let _, ((env_params, ctx_params), userimpls) = interp_context_evars env0 evdref paramsl in + let ctx_params = List.map (fun d -> map_rel_decl EConstr.Unsafe.to_constr d) ctx_params in let indnames = List.map (fun ind -> ind.ind_name) indl in (* Names of parameters as arguments of the inductive type (defs removed) *) @@ -850,16 +852,16 @@ let interp_fix_context env evdref isfix fix = let interp_fix_ccl evdref impls (env,_) fix = let (c, impl) = interp_type_evars_impls ~impls env evdref fix.fix_type in - (EConstr.Unsafe.to_constr c, impl) + (c, impl) let interp_fix_body env_rec evdref impls (_,ctx) fix ccl = + let open EConstr in Option.map (fun body -> let env = push_rel_context ctx env_rec in let body = interp_casted_constr_evars env evdref ~impls body ccl in - let body = EConstr.Unsafe.to_constr body in it_mkLambda_or_LetIn body ctx) fix.fix_body -let build_fix_type (_,ctx) ccl = Term.it_mkProd_or_LetIn ccl ctx +let build_fix_type (_,ctx) ccl = EConstr.it_mkProd_or_LetIn ccl ctx let declare_fix ?(opaque = false) (_,poly,_ as kind) pl ctx f ((def,_),eff) t imps = let ce = definition_entry ~opaque ~types:t ~poly ~univs:ctx ~eff def in @@ -899,7 +901,7 @@ let fixsub_module = subtac_dir @ ["Wf"] let tactics_module = subtac_dir @ ["Tactics"] let init_reference dir s () = Coqlib.gen_reference "Command" dir s -let init_constant dir s () = Coqlib.gen_constant "Command" dir s +let init_constant dir s () = EConstr.of_constr (Coqlib.gen_constant "Command" dir s) let make_ref l s = init_reference l s let fix_proto = init_constant tactics_module "fix_proto" let fix_sub_ref = make_ref fixsub_module "Fix_sub" @@ -907,14 +909,17 @@ let measure_on_R_ref = make_ref fixsub_module "MR" let well_founded = init_constant ["Init"; "Wf"] "well_founded" let mkSubset name typ prop = let open EConstr in - EConstr.Unsafe.to_constr (mkApp (EConstr.of_constr (Universes.constr_of_global (delayed_force build_sigma).typ), - [| typ; mkLambda (name, typ, prop) |])) + mkApp (EConstr.of_constr (Universes.constr_of_global (delayed_force build_sigma).typ), + [| typ; mkLambda (name, typ, prop) |]) let sigT = Lazy.from_fun build_sigma_type let make_qref s = Qualid (Loc.ghost, qualid_of_string s) let lt_ref = make_qref "Init.Peano.lt" -let rec telescope = function +let rec telescope l = + let open EConstr in + let open Vars in + match l with | [] -> assert false | [LocalAssum (n, t)] -> t, [LocalDef (n, mkRel 1, t)], mkRel 1 | LocalAssum (n, t) :: tl -> @@ -924,7 +929,9 @@ let rec telescope = function let t = RelDecl.get_type decl in let pred = mkLambda (RelDecl.get_name decl, t, ty) in let ty = Universes.constr_of_global (Lazy.force sigT).typ in + let ty = EConstr.of_constr ty in let intro = Universes.constr_of_global (Lazy.force sigT).intro in + let intro = EConstr.of_constr intro in let sigty = mkApp (ty, [|t; pred|]) in let intro = mkApp (intro, [|lift k t; lift k pred; mkRel k; constr|]) in (sigty, pred :: tys, (succ k, intro))) @@ -934,9 +941,11 @@ let rec telescope = function (fun pred decl (prev, subst) -> let t = RelDecl.get_type decl in let p1 = Universes.constr_of_global (Lazy.force sigT).proj1 in + let p1 = EConstr.of_constr p1 in let p2 = Universes.constr_of_global (Lazy.force sigT).proj2 in - let proj1 = applistc p1 [t; pred; prev] in - let proj2 = applistc p2 [t; pred; prev] in + let p2 = EConstr.of_constr p2 in + let proj1 = applist (p1, [t; pred; prev]) in + let proj2 = applist (p2, [t; pred; prev]) in (lift 1 proj2, LocalDef (get_name decl, proj1, t) :: subst)) (List.rev tys) tl (mkRel 1, []) in ty, (LocalDef (n, last, t) :: subst), constr @@ -945,11 +954,12 @@ let rec telescope = function ty, (LocalDef (n, b, t) :: subst), lift 1 term let nf_evar_context sigma ctx = - List.map (map_constr (fun c -> EConstr.Unsafe.to_constr (Evarutil.nf_evar sigma (EConstr.of_constr c)))) ctx + List.map (map_constr (fun c -> Evarutil.nf_evar sigma c)) ctx let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation = let open EConstr in let open Vars in + let lift_rel_context n l = Termops.map_rel_context_with_binders (liftn n) l in Coqlib.check_required_library ["Coq";"Program";"Wf"]; let env = Global.env() in let ctx = Evd.make_evar_universe_context env pl in @@ -960,10 +970,8 @@ let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation = let top_arity = interp_type_evars top_env evdref arityc in let full_arity = it_mkProd_or_LetIn top_arity binders_rel in let argtyp, letbinders, make = telescope binders_rel in - let make = EConstr.of_constr make in let argname = Id.of_string "recarg" in let arg = LocalAssum (Name argname, argtyp) in - let argtyp = EConstr.of_constr argtyp in let binders = letbinders @ [arg] in let binders_env = push_rel_context binders_rel env in let rel, _ = interp_constr_evars_impls env evdref r in @@ -978,10 +986,11 @@ let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation = let ctx, ar = Reductionops.splay_prod_n env !evdref 2 relty in match ctx, EConstr.kind !evdref ar with | [LocalAssum (_,t); LocalAssum (_,u)], Sort (Prop Null) - when Reductionops.is_conv env !evdref (EConstr.of_constr t) (EConstr.of_constr u) -> t + when Reductionops.is_conv env !evdref t u -> t | _, _ -> error () with e when CErrors.noncritical e -> error () in + let relargty = EConstr.Unsafe.to_constr relargty in let measure = interp_casted_constr_evars binders_env evdref measure relargty in let wf_rel, wf_rel_fun, measure_fn = let measure_body, measure = @@ -996,7 +1005,7 @@ let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation = subst1 y measure_body |]) in wf_rel, wf_rel_fun, measure in - let wf_proof = mkApp (EConstr.of_constr (delayed_force well_founded), [| argtyp ; wf_rel |]) in + let wf_proof = mkApp (delayed_force well_founded, [| argtyp ; wf_rel |]) in let argid' = Id.of_string (Id.to_string argname ^ "'") in let wfarg len = LocalAssum (Name argid', mkSubset (Name argid') argtyp @@ -1014,7 +1023,6 @@ let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation = (* substitute the projection of wfarg for something, now intern_arity is in wfarg :: arg *) let intern_fun_arity_prod = it_mkProd_or_LetIn intern_arity [wfarg 1] in - let intern_fun_arity_prod = EConstr.Unsafe.to_constr intern_fun_arity_prod in let intern_fun_binder = LocalAssum (Name (add_suffix recname "'"), intern_fun_arity_prod) in let curry_fun = let wfpred = mkLambda (Name argid', argtyp, wf_rel_fun (mkRel 1) (mkRel (2 * len + 4))) in @@ -1022,16 +1030,13 @@ let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation = let arg = mkApp (intro, [| argtyp; wfpred; lift 1 make; mkRel 1 |]) in let app = mkApp (mkRel (2 * len + 2 (* recproof + orig binders + current binders *)), [| arg |]) in let rcurry = mkApp (rel, [| measure; lift len measure |]) in - let rcurry = EConstr.Unsafe.to_constr rcurry in let lam = LocalAssum (Name (Id.of_string "recproof"), rcurry) in let body = it_mkLambda_or_LetIn app (lam :: binders_rel) in let ty = it_mkProd_or_LetIn (lift 1 top_arity) (lam :: binders_rel) in - let body = EConstr.Unsafe.to_constr body in - let ty = EConstr.Unsafe.to_constr ty in LocalDef (Name recname, body, ty) in let fun_bl = intern_fun_binder :: [arg] in - let lift_lets = Termops.lift_rel_context 1 letbinders in + let lift_lets = lift_rel_context 1 letbinders in let intern_body = let ctx = LocalAssum (Name recname, get_type curry_fun) :: binders_rel in let (r, l, impls, scopes) = @@ -1099,6 +1104,7 @@ let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation = let interp_recursive isfix fixl notations = let open Context.Named.Declaration in + let open EConstr in let env = Global.env() in let fixnames = List.map (fun fix -> fix.fix_name) fixl in @@ -1119,20 +1125,19 @@ let interp_recursive isfix fixl notations = let fixctximpenvs, fixctximps = List.split fiximppairs in let fixccls,fixcclimps = List.split (List.map3 (interp_fix_ccl evdref) fixctximpenvs fixctxs fixl) in let fixtypes = List.map2 build_fix_type fixctxs fixccls in - let fixtypes = List.map (fun c -> EConstr.Unsafe.to_constr (nf_evar !evdref (EConstr.of_constr c))) fixtypes in + let fixtypes = List.map (fun c -> nf_evar !evdref c) fixtypes in let fiximps = List.map3 (fun ctximps cclimps (_,ctx) -> ctximps@(Impargs.lift_implicits (List.length ctx) cclimps)) fixctximps fixcclimps fixctxs in let rec_sign = List.fold_left2 (fun env' id t -> - if Flags.is_program_mode () then - let sort = Evarutil.evd_comb1 (Typing.type_of ~refresh:true env) evdref (EConstr.of_constr t) in - let sort = EConstr.Unsafe.to_constr sort in + if Flags.is_program_mode () then + let sort = Evarutil.evd_comb1 (Typing.type_of ~refresh:true env) evdref t in let fixprot = try let app = mkApp (delayed_force fix_proto, [|sort; t|]) in - EConstr.Unsafe.to_constr (Typing.e_solve_evars env evdref (EConstr.of_constr app)) + Typing.e_solve_evars env evdref app with e when CErrors.noncritical e -> t in LocalAssum (id,fixprot) :: env' @@ -1142,6 +1147,8 @@ let interp_recursive isfix fixl notations = let env_rec = push_named_context rec_sign env in (* Get interpretation metadatas *) + let fixtypes = List.map EConstr.Unsafe.to_constr fixtypes in + let fixccls = List.map EConstr.Unsafe.to_constr fixccls in let impls = compute_internalization_env env Recursive fixnames fixtypes fiximps in (* Interp bodies with rollback because temp use of notations/implicit *) @@ -1156,6 +1163,7 @@ let interp_recursive isfix fixl notations = (* Instantiate evars and check all are resolved *) let evd = consider_remaining_unif_problems env_rec !evdref in let evd, nf = nf_evars_and_universes evd in + let fixdefs = List.map (fun c -> Option.map EConstr.Unsafe.to_constr c) fixdefs in let fixdefs = List.map (Option.map nf) fixdefs in let fixtypes = List.map nf fixtypes in let fixctxnames = List.map (fun (_,ctx) -> List.map RelDecl.get_name ctx) fixctxs in diff --git a/toplevel/himsg.ml b/toplevel/himsg.ml index 851b87903..b689dfcf9 100644 --- a/toplevel/himsg.ml +++ b/toplevel/himsg.ml @@ -29,56 +29,56 @@ module RelDecl = Context.Rel.Declaration (* This simplifies the typing context of Cases clauses *) (* hope it does not disturb other typing contexts *) -let contract env lc = +let contract env sigma lc = + let open EConstr in let l = ref [] in let contract_context decl env = match decl with - | LocalDef (_,c',_) when isRel c' -> + | LocalDef (_,c',_) when isRel sigma c' -> l := (Vars.substl !l c') :: !l; env | _ -> let t = Vars.substl !l (RelDecl.get_type decl) in - let decl = decl |> RelDecl.map_name (named_hd env t) |> RelDecl.map_value (Vars.substl !l) |> RelDecl.set_type t in + let decl = decl |> RelDecl.map_name (named_hd env (EConstr.Unsafe.to_constr t)) |> RelDecl.map_value (Vars.substl !l) |> RelDecl.set_type t in l := (mkRel 1) :: List.map (Vars.lift 1) !l; push_rel decl env in let env = process_rel_context contract_context env in (env, List.map (Vars.substl !l) lc) -let contract2 env a b = match contract env [a;b] with +let contract2 env sigma a b = match contract env sigma [a;b] with | env, [a;b] -> env,a,b | _ -> assert false -let contract3 env a b c = match contract env [a;b;c] with +let contract3 env sigma a b c = match contract env sigma [a;b;c] with | env, [a;b;c] -> env,a,b,c | _ -> assert false -let contract4 env a b c d = match contract env [a;b;c;d] with +let contract4 env sigma a b c d = match contract env sigma [a;b;c;d] with | env, [a;b;c;d] -> (env,a,b,c),d | _ -> assert false -let contract1_vect env a v = - match contract env (a :: Array.to_list v) with +let contract1_vect env sigma a v = + match contract env sigma (a :: Array.to_list v) with | env, a::l -> env,a,Array.of_list l | _ -> assert false -let rec contract3' env a b c = function +let rec contract3' env sigma a b c = function | OccurCheck (evk,d) -> - let d = EConstr.Unsafe.to_constr d in - let x,d = contract4 env a b c d in x,OccurCheck(evk, EConstr.of_constr d) + let x,d = contract4 env sigma a b c d in x,OccurCheck(evk, d) | NotClean ((evk,args),env',d) -> - let d = EConstr.Unsafe.to_constr d in - let args = Array.map EConstr.Unsafe.to_constr args in - let env',d,args = contract1_vect env' d args in - contract3 env a b c,NotClean((evk,Array.map EConstr.of_constr args),env',EConstr.of_constr d) + let env',d,args = contract1_vect env' sigma d args in + contract3 env sigma a b c,NotClean((evk,args),env',d) | ConversionFailed (env',t1,t2) -> - let t1 = EConstr.Unsafe.to_constr t1 in - let t2 = EConstr.Unsafe.to_constr t2 in - let (env',t1,t2) = contract2 env' t1 t2 in - contract3 env a b c, ConversionFailed (env',EConstr.of_constr t1,EConstr.of_constr t2) + let (env',t1,t2) = contract2 env' sigma t1 t2 in + contract3 env sigma a b c, ConversionFailed (env',t1,t2) | NotSameArgSize | NotSameHead | NoCanonicalStructure | MetaOccurInBody _ | InstanceNotSameType _ | ProblemBeyondCapabilities - | UnifUnivInconsistency _ as x -> contract3 env a b c, x + | UnifUnivInconsistency _ as x -> contract3 env sigma a b c, x | CannotSolveConstraint ((pb,env',t,u),x) -> - let env',t,u = contract2 env' t u in - let y,x = contract3' env a b c x in + let t = EConstr.of_constr t in + let u = EConstr.of_constr u in + let env',t,u = contract2 env' sigma t u in + let t = EConstr.Unsafe.to_constr t in + let u = EConstr.Unsafe.to_constr u in + let y,x = contract3' env sigma a b c x in y,CannotSolveConstraint ((pb,env',t,u),x) (** Ad-hoc reductions *) @@ -820,29 +820,19 @@ let explain_pretype_error env sigma err = match err with | CantFindCaseType c -> explain_cant_find_case_type env sigma c | ActualTypeNotCoercible (j,t,e) -> - let j = to_unsafe_judgment j in - let t = EConstr.Unsafe.to_constr t in let {uj_val = c; uj_type = actty} = j in - let (env, c, actty, expty), e = contract3' env c actty t e in + let (env, c, actty, expty), e = contract3' env sigma c actty t e in let j = {uj_val = c; uj_type = actty} in - explain_actual_type env sigma (on_judgment EConstr.of_constr j) (EConstr.of_constr expty) (Some e) + explain_actual_type env sigma j expty (Some e) | UnifOccurCheck (ev,rhs) -> explain_occur_check env sigma ev rhs | UnsolvableImplicit (evk,exp) -> explain_unsolvable_implicit env sigma evk exp | VarNotFound id -> explain_var_not_found env id | UnexpectedType (actual,expect) -> - let actual = EConstr.Unsafe.to_constr actual in - let expect = EConstr.Unsafe.to_constr expect in - let env, actual, expect = contract2 env actual expect in - let actual = EConstr.of_constr actual in - let expect = EConstr.of_constr expect in + let env, actual, expect = contract2 env sigma actual expect in explain_unexpected_type env sigma actual expect | NotProduct c -> explain_not_product env sigma c | CannotUnify (m,n,e) -> - let m = EConstr.Unsafe.to_constr m in - let n = EConstr.Unsafe.to_constr n in - let env, m, n = contract2 env m n in - let m = EConstr.of_constr m in - let n = EConstr.of_constr n in + let env, m, n = contract2 env sigma m n in explain_cannot_unify env sigma m n e | CannotUnifyLocal (m,n,sn) -> explain_cannot_unify_local env sigma m n sn | CannotGeneralize ty -> explain_refiner_cannot_generalize env sigma ty diff --git a/toplevel/record.ml b/toplevel/record.ml index ec719bca8..af2377e51 100644 --- a/toplevel/record.ml +++ b/toplevel/record.ml @@ -114,6 +114,7 @@ let typecheck_params_and_fields def id pl t ps nots fs = | LocalPattern _ -> assert false) ps in let impls_env, ((env1,newps), imps) = interp_context_evars env0 evars ps in + let newps = List.map (fun d -> Termops.map_rel_decl EConstr.Unsafe.to_constr d) newps in let t', template = match t with | Some t -> let env = push_rel_context newps env0 in diff --git a/toplevel/vernacentries.ml b/toplevel/vernacentries.ml index cf3da7e02..f3d73b76c 100644 --- a/toplevel/vernacentries.ml +++ b/toplevel/vernacentries.ml @@ -1651,7 +1651,7 @@ let print_about_hyp_globs ref_or_by_not glnumopt = let natureofid = match decl with | LocalAssum _ -> "Hypothesis" | LocalDef (_,bdy,_) ->"Constant (let in)" in - v 0 (pr_id id ++ str":" ++ pr_constr (NamedDecl.get_type decl) ++ fnl() ++ fnl() + v 0 (pr_id id ++ str":" ++ pr_econstr (NamedDecl.get_type decl) ++ fnl() ++ fnl() ++ str natureofid ++ str " of the goal context.") with (* fallback to globals *) | NoHyp | Not_found -> print_about ref_or_by_not -- cgit v1.2.3 From 4e9cebb0641927f11a21cbb50828974f910cfe47 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 29 Nov 2016 15:25:06 +0100 Subject: Putting back the subst_defined_metas_evars function in the old term API. It seems this is a performance-critical function for unification-heavy code. In particular, tactics relying on meta unification suffered an important penalty after this function was rewritten with the evar-insensitive API, as witnessed e.g. by Ncring_polynom whose compilation time increased by ~30%. I am not sure about the specification of this function, but it seems safer to revert the changes and just do it the old way. It may even disappear if we get rid of the old unification algorithm at some point. --- pretyping/unification.ml | 18 +++++++++++------- 1 file changed, 11 insertions(+), 7 deletions(-) diff --git a/pretyping/unification.ml b/pretyping/unification.ml index 0d6dcffc1..589201fe2 100644 --- a/pretyping/unification.ml +++ b/pretyping/unification.ml @@ -601,16 +601,20 @@ let isAllowedEvar sigma flags c = match EConstr.kind sigma c with let subst_defined_metas_evars sigma (bl,el) c = - let rec substrec c = match EConstr.kind sigma c with + (** This seems to be performance-critical, and using the evar-insensitive + primitives blow up the time passed in this function. *) + let c = EConstr.Unsafe.to_constr c in + let rec substrec c = match kind_of_term c with | Meta i -> let select (j,_,_) = Int.equal i j in - substrec (pi2 (List.find select bl)) + substrec (EConstr.Unsafe.to_constr (pi2 (List.find select bl))) | Evar (evk,args) -> - let select (_,(evk',args'),_) = Evar.equal evk evk' && Array.equal (EConstr.eq_constr sigma) args args' in - (try substrec (pi3 (List.find select el)) - with Not_found -> EConstr.map sigma substrec c) - | _ -> EConstr.map sigma substrec c - in try Some (substrec c) with Not_found -> None + let eq c1 c2 = Constr.equal c1 (EConstr.Unsafe.to_constr c2) in + let select (_,(evk',args'),_) = Evar.equal evk evk' && Array.for_all2 eq args args' in + (try substrec (EConstr.Unsafe.to_constr (pi3 (List.find select el))) + with Not_found -> Constr.map substrec c) + | _ -> Constr.map substrec c + in try Some (EConstr.of_constr (substrec c)) with Not_found -> None let check_compatibility env pbty flags (sigma,metasubst,evarsubst : subst0) tyM tyN = match subst_defined_metas_evars sigma (metasubst,[]) tyM with -- cgit v1.2.3 From d549d9d3d169fbfc5f555e3e4f22f46301161d53 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 29 Nov 2016 16:30:00 +0100 Subject: Do not ask for a normalized goal to get hypotheses and conclusions. This is now useless as this returns evar-constrs, so that all functions acting on them should be insensitive to evar-normalization. --- engine/proofview.ml | 3 --- engine/proofview.mli | 8 ++------ proofs/tacmach.mli | 10 +++++----- tactics/hipattern.mli | 6 +++--- tactics/tacticals.mli | 2 +- tactics/tactics.ml | 12 ++++++------ 6 files changed, 17 insertions(+), 24 deletions(-) diff --git a/engine/proofview.ml b/engine/proofview.ml index 0a18cf191..71e9acc88 100644 --- a/engine/proofview.ml +++ b/engine/proofview.ml @@ -999,9 +999,6 @@ module Goal = struct let concl { concl=concl } = concl let extra { sigma=sigma; self=self } = goal_extra sigma self - let raw_concl { concl=concl } = concl - - let gmake_with info env sigma goal = { env = Environ.reset_with_named_context (Evd.evar_filtered_hyps info) env ; sigma = sigma ; diff --git a/engine/proofview.mli b/engine/proofview.mli index 025e3de20..4f662b294 100644 --- a/engine/proofview.mli +++ b/engine/proofview.mli @@ -485,16 +485,12 @@ module Goal : sig 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 -> constr - val hyps : ([ `NF ], 'r) t -> named_context + val concl : ('a, 'r) t -> constr + val hyps : ('a, 'r) t -> named_context 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 -> constr - type ('a, 'b) enter = { enter : 'r. ('a, 'r) t -> 'b } diff --git a/proofs/tacmach.mli b/proofs/tacmach.mli index 3b23a6ab4..1992cec65 100644 --- a/proofs/tacmach.mli +++ b/proofs/tacmach.mli @@ -107,19 +107,19 @@ module New : sig val project : ('a, 'r) Proofview.Goal.t -> Evd.evar_map val pf_env : ('a, 'r) Proofview.Goal.t -> Environ.env - val pf_concl : ([ `NF ], 'r) Proofview.Goal.t -> types + val pf_concl : ('a, 'r) Proofview.Goal.t -> types val pf_unsafe_type_of : ('a, 'r) Proofview.Goal.t -> constr -> types val pf_type_of : ('a, 'r) Proofview.Goal.t -> constr -> evar_map * types val pf_conv_x : ('a, 'r) Proofview.Goal.t -> t -> t -> bool - val pf_get_new_id : identifier -> ([ `NF ], 'r) Proofview.Goal.t -> identifier + val pf_get_new_id : identifier -> ('a, 'r) Proofview.Goal.t -> identifier val pf_ids_of_hyps : ('a, 'r) Proofview.Goal.t -> identifier list val pf_hyps_types : ('a, 'r) Proofview.Goal.t -> (identifier * types) list - val pf_get_hyp : identifier -> ([ `NF ], 'r) Proofview.Goal.t -> named_declaration - val pf_get_hyp_typ : identifier -> ([ `NF ], 'r) Proofview.Goal.t -> types - val pf_last_hyp : ([ `NF ], 'r) Proofview.Goal.t -> named_declaration + val pf_get_hyp : identifier -> ('a, 'r) Proofview.Goal.t -> named_declaration + val pf_get_hyp_typ : identifier -> ('a, 'r) Proofview.Goal.t -> types + val pf_last_hyp : ('a, 'r) Proofview.Goal.t -> named_declaration val pf_nf_concl : ([ `LZ ], 'r) Proofview.Goal.t -> types val pf_reduce_to_quantified_ind : ('a, 'r) Proofview.Goal.t -> types -> pinductive * types diff --git a/tactics/hipattern.mli b/tactics/hipattern.mli index 65ba0aad0..c46817f50 100644 --- a/tactics/hipattern.mli +++ b/tactics/hipattern.mli @@ -121,11 +121,11 @@ val match_with_equation: (** Match terms [eq A t u], [identity A t u] or [JMeq A t A u] Returns associated lemmas and [A,t,u] or fails PatternMatchingFailure *) -val find_eq_data_decompose : ([ `NF ], 'r) Proofview.Goal.t -> constr -> +val find_eq_data_decompose : ('a, 'r) Proofview.Goal.t -> constr -> coq_eq_data * Univ.universe_instance * (types * constr * constr) (** Idem but fails with an error message instead of PatternMatchingFailure *) -val find_this_eq_data_decompose : ([ `NF ], 'r) Proofview.Goal.t -> constr -> +val find_this_eq_data_decompose : ('a, 'r) Proofview.Goal.t -> constr -> coq_eq_data * Univ.universe_instance * (types * constr * constr) (** A variant that returns more informative structure on the equality found *) @@ -146,7 +146,7 @@ val is_matching_sigma : evar_map -> constr -> bool val match_eqdec : evar_map -> constr -> bool * constr * constr * constr * constr (** Match an equality up to conversion; returns [(eq,t1,t2)] in normal form *) -val dest_nf_eq : ([ `NF ], 'r) Proofview.Goal.t -> constr -> (constr * constr * constr) +val dest_nf_eq : ('a, 'r) Proofview.Goal.t -> constr -> (constr * constr * constr) (** Match a negation *) val is_matching_not : evar_map -> constr -> bool diff --git a/tactics/tacticals.mli b/tactics/tacticals.mli index e9f623100..4bb745875 100644 --- a/tactics/tacticals.mli +++ b/tactics/tacticals.mli @@ -229,7 +229,7 @@ module New : sig val tclTIMEOUT : int -> unit tactic -> unit tactic val tclTIME : string option -> 'a tactic -> 'a tactic - val nLastDecls : ([ `NF ], 'r) Proofview.Goal.t -> int -> named_context + val nLastDecls : ('a, 'r) Proofview.Goal.t -> int -> named_context val ifOnHyp : (identifier * types -> bool) -> (identifier -> unit Proofview.tactic) -> (identifier -> unit Proofview.tactic) -> diff --git a/tactics/tactics.ml b/tactics/tactics.ml index 4bf848a9c..de3572155 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -208,7 +208,7 @@ let convert_concl ?(check=true) ty k = Proofview.Goal.enter { enter = begin fun gl -> let env = Proofview.Goal.env gl in let store = Proofview.Goal.extra gl in - let conclty = Proofview.Goal.raw_concl gl in + let conclty = Proofview.Goal.concl gl in Refine.refine ~unsafe:true { run = begin fun sigma -> let Sigma ((), sigma, p) = if check then begin @@ -228,7 +228,7 @@ let convert_hyp ?(check=true) d = Proofview.Goal.enter { enter = begin fun gl -> let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in - let ty = Proofview.Goal.raw_concl gl in + let ty = Proofview.Goal.concl gl in let store = Proofview.Goal.extra gl in let sign = convert_hyp check (named_context_val env) sigma d in let env = reset_with_named_context sign env in @@ -328,7 +328,7 @@ let move_hyp id dest = Proofview.Goal.enter { enter = begin fun gl -> let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in - let ty = Proofview.Goal.raw_concl gl in + let ty = Proofview.Goal.concl gl in let store = Proofview.Goal.extra gl in let sign = named_context_val env in let sign' = move_hyp_in_named_context sigma id dest sign in @@ -756,7 +756,7 @@ let e_reduct_option ?(check=false) redfun = function let e_change_in_concl (redfun,sty) = Proofview.Goal.s_enter { s_enter = begin fun gl -> let sigma = Proofview.Goal.sigma gl in - let Sigma (c, sigma, p) = redfun.e_redfun (Proofview.Goal.env gl) sigma (Proofview.Goal.raw_concl gl) in + let Sigma (c, sigma, p) = redfun.e_redfun (Proofview.Goal.env gl) sigma (Proofview.Goal.concl gl) in Sigma (convert_concl_no_check c sty, sigma, p) end } @@ -4340,7 +4340,7 @@ let pose_induction_arg_then isrec with_evars (is_arg_pure_hyp,from_prefix) elim Proofview.Goal.s_enter { s_enter = begin fun gl -> let sigma = Proofview.Goal.sigma gl in let env = Proofview.Goal.env gl in - let ccl = Proofview.Goal.raw_concl gl in + let ccl = Proofview.Goal.concl gl in let store = Proofview.Goal.extra gl in let check = check_enough_applied env sigma elim in let Sigma (c, sigma', p) = use_bindings env sigma elim false (c0,lbind) t0 in @@ -4409,7 +4409,7 @@ let induction_gen clear_flag isrec with_evars elim let env = Proofview.Goal.env gl in let sigma = Proofview.Goal.sigma gl in let evd = Sigma.to_evar_map sigma in - let ccl = Proofview.Goal.raw_concl gl in + let ccl = Proofview.Goal.concl gl in let cls = Option.default allHypsAndConcl cls in let t = typ_of env sigma c in let is_arg_pure_hyp = -- cgit v1.2.3 From 390fd4ac0a969103caeb5db3e5138e26f9a533de Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 29 Nov 2016 17:49:11 +0100 Subject: Chasing a few unsafe constr coercions. --- engine/evd.ml | 29 ----------------------------- engine/evd.mli | 13 ------------- pretyping/evarconv.ml | 8 +++++++- pretyping/evarsolve.ml | 8 +++++++- proofs/tacmach.ml | 6 ++---- tactics/btermdn.ml | 33 +++++++++++++++++---------------- tactics/btermdn.mli | 2 +- tactics/equality.ml | 13 ------------- tactics/hints.ml | 20 +++++++++++++++----- tactics/tactics.ml | 7 +++++-- toplevel/himsg.ml | 28 +++++++++++----------------- 11 files changed, 65 insertions(+), 102 deletions(-) diff --git a/engine/evd.ml b/engine/evd.ml index 6380e3fc1..b30702d0d 100644 --- a/engine/evd.ml +++ b/engine/evd.ml @@ -941,37 +941,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 *) diff --git a/engine/evd.mli b/engine/evd.mli index 699b52c2b..eeb9fd861 100644 --- a/engine/evd.mli +++ b/engine/evd.mli @@ -569,19 +569,6 @@ 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. *) diff --git a/pretyping/evarconv.ml b/pretyping/evarconv.ml index f5cab070e..1cbea68dd 100644 --- a/pretyping/evarconv.ml +++ b/pretyping/evarconv.ml @@ -973,7 +973,13 @@ let apply_on_subterm env evdref f c t = (* By using eq_constr, we make an approximation, for instance, we *) (* could also be interested in finding a term u convertible to t *) (* such that c occurs in u *) - if e_eq_constr_univs evdref (EConstr.Unsafe.to_constr c) (EConstr.Unsafe.to_constr t) then f k + let eq_constr c1 c2 = match EConstr.eq_constr_universes !evdref c1 c2 with + | None -> false + | Some cstr -> + try ignore (Evd.add_universe_constraints !evdref cstr); true + with UniversesDiffer -> false + in + if eq_constr c t then f k else match EConstr.kind !evdref t with | Evar (evk,args) -> diff --git a/pretyping/evarsolve.ml b/pretyping/evarsolve.ml index ff4736528..28e63d04b 100644 --- a/pretyping/evarsolve.ml +++ b/pretyping/evarsolve.ml @@ -1263,7 +1263,13 @@ type conv_fun_bool = let solve_refl ?(can_drop=false) conv_algo env evd pbty evk argsv1 argsv2 = let evdref = ref evd in - if Array.equal (fun c1 c2 -> e_eq_constr_univs evdref (EConstr.Unsafe.to_constr c1) (EConstr.Unsafe.to_constr c2) ) argsv1 argsv2 then !evdref else + let eq_constr c1 c2 = match EConstr.eq_constr_universes !evdref c1 c2 with + | None -> false + | Some cstr -> + try ignore (Evd.add_universe_constraints !evdref cstr); true + with UniversesDiffer -> false + in + if Array.equal eq_constr argsv1 argsv2 then !evdref else (* Filter and restrict if needed *) let args = Array.map2 (fun a1 a2 -> (a1, a2)) argsv1 argsv2 in let untypedfilter = diff --git a/proofs/tacmach.ml b/proofs/tacmach.ml index 4b027a127..97c5cda77 100644 --- a/proofs/tacmach.ml +++ b/proofs/tacmach.ml @@ -41,10 +41,8 @@ let project = Refiner.project let pf_env = Refiner.pf_env let pf_hyps = Refiner.pf_hyps -let test_conversion cb env sigma c1 c2 = - let c1 = EConstr.Unsafe.to_constr c1 in - let c2 = EConstr.Unsafe.to_constr c2 in - test_conversion cb env sigma c1 c2 +let test_conversion env sigma pb c1 c2 = + Reductionops.check_conv ~pb env sigma c1 c2 let pf_concl gls = Goal.V82.concl (project gls) (sig_it gls) let pf_hyps_types gls = diff --git a/tactics/btermdn.ml b/tactics/btermdn.ml index 491bc8b4a..b4a235ba8 100644 --- a/tactics/btermdn.ml +++ b/tactics/btermdn.ml @@ -8,6 +8,7 @@ open Util open Term +open EConstr open Names open Pattern open Globnames @@ -38,18 +39,18 @@ let decomp_pat = in decrec [] -let decomp = - let rec decrec acc c = match kind_of_term c with +let decomp sigma t = + let rec decrec acc c = match EConstr.kind sigma c with | App (f,l) -> decrec (Array.fold_right (fun a l -> a::l) l acc) f | Proj (p, c) -> (mkConst (Projection.constant p), c :: acc) | Cast (c1,_,_) -> decrec acc c1 | _ -> (c,acc) in - decrec [] + decrec [] t -let constr_val_discr t = - let c, l = decomp t in - match kind_of_term c with +let constr_val_discr sigma t = + let c, l = decomp sigma t in + match EConstr.kind sigma c with | Ind (ind_sp,u) -> Label(GRLabel (IndRef ind_sp),l) | Construct (cstr_sp,u) -> Label(GRLabel (ConstructRef cstr_sp),l) | Var id -> Label(GRLabel (VarRef id),l) @@ -66,9 +67,9 @@ let constr_pat_discr t = | PRef ((VarRef v) as ref), args -> Some(GRLabel ref,args) | _ -> None -let constr_val_discr_st (idpred,cpred) t = - let c, l = decomp t in - match kind_of_term c with +let constr_val_discr_st sigma (idpred,cpred) t = + let c, l = decomp sigma t in + match EConstr.kind sigma c with | Const (c,u) -> if Cpred.mem c cpred then Everything else Label(GRLabel (ConstRef c),l) | Ind (ind_sp,u) -> Label(GRLabel (IndRef ind_sp),l) | Construct (cstr_sp,u) -> Label(GRLabel (ConstructRef cstr_sp),l) @@ -105,11 +106,11 @@ let bounded_constr_pat_discr_st st (t,depth) = | None -> None | Some (c,l) -> Some(c,List.map (fun c -> (c,depth-1)) l) -let bounded_constr_val_discr_st st (t,depth) = +let bounded_constr_val_discr_st sigma st (t,depth) = if Int.equal depth 0 then Nothing else - match constr_val_discr_st st t with + match constr_val_discr_st sigma st t with | Label (c,l) -> Label(c,List.map (fun c -> (c,depth-1)) l) | Nothing -> Nothing | Everything -> Everything @@ -122,11 +123,11 @@ let bounded_constr_pat_discr (t,depth) = | None -> None | Some (c,l) -> Some(c,List.map (fun c -> (c,depth-1)) l) -let bounded_constr_val_discr (t,depth) = +let bounded_constr_val_discr sigma (t,depth) = if Int.equal depth 0 then Nothing else - match constr_val_discr t with + match constr_val_discr sigma t with | Label (c,l) -> Label(c,List.map (fun c -> (c,depth-1)) l) | Nothing -> Nothing | Everything -> Everything @@ -162,13 +163,13 @@ struct (fun dn (c,v) -> Dn.rmv dn (bounded_constr_pat_discr_st st) ((c,!dnet_depth),v)) - let lookup = function + let lookup sigma = function | None -> (fun dn t -> - Dn.lookup dn bounded_constr_val_discr (t,!dnet_depth)) + Dn.lookup dn (bounded_constr_val_discr sigma) (t,!dnet_depth)) | Some st -> (fun dn t -> - Dn.lookup dn (bounded_constr_val_discr_st st) (t,!dnet_depth)) + Dn.lookup dn (bounded_constr_val_discr_st sigma st) (t,!dnet_depth)) let app f dn = Dn.app f dn diff --git a/tactics/btermdn.mli b/tactics/btermdn.mli index 8ca5549b8..2a5e7c345 100644 --- a/tactics/btermdn.mli +++ b/tactics/btermdn.mli @@ -33,7 +33,7 @@ sig val add : transparent_state option -> t -> (constr_pattern * Z.t) -> t val rmv : transparent_state option -> t -> (constr_pattern * Z.t) -> t - val lookup : transparent_state option -> t -> constr -> Z.t list + val lookup : Evd.evar_map -> transparent_state option -> t -> EConstr.constr -> Z.t list val app : (Z.t -> unit) -> t -> unit end diff --git a/tactics/equality.ml b/tactics/equality.ml index 122b64777..7f7a07b8f 100644 --- a/tactics/equality.ml +++ b/tactics/equality.ml @@ -1300,19 +1300,6 @@ let build_injector env sigma dflt c cpath = let sigma, (injcode,resty,_) = build_injrec env sigma dflt c cpath in sigma, (injcode,resty) -(* -let try_delta_expand env sigma t = - let whdt = whd_all env sigma t in - let rec hd_rec c = - match kind_of_term c with - | Construct _ -> whdt - | App (f,_) -> hd_rec f - | Cast (c,_,_) -> hd_rec c - | _ -> t - in - hd_rec whdt -*) - let eq_dec_scheme_kind_name = ref (fun _ -> failwith "eq_dec_scheme undefined") let set_eq_dec_scheme_kind k = eq_dec_scheme_kind_name := (fun _ -> k) diff --git a/tactics/hints.ml b/tactics/hints.ml index ffd19ac6e..17c81064d 100644 --- a/tactics/hints.ml +++ b/tactics/hints.ml @@ -250,9 +250,8 @@ let rebuild_dn st se = in { se with sentry_bnet = dn' } -let lookup_tacs concl st se = - let concl = EConstr.Unsafe.to_constr concl in - let l' = Bounded_net.lookup st se.sentry_bnet concl in +let lookup_tacs sigma concl st se = + let l' = Bounded_net.lookup sigma st se.sentry_bnet concl in let sl' = List.stable_sort pri_order_int l' in List.merge pri_order_int se.sentry_nopat sl' @@ -510,6 +509,17 @@ struct (** Warn about no longer typable hint? *) None + let head_evar sigma c = + let rec hrec c = match EConstr.kind sigma c with + | Evar (evk,_) -> evk + | Case (_,_,c,_) -> hrec c + | App (c,_) -> hrec c + | Cast (c,_,_) -> hrec c + | Proj (p, c) -> hrec c + | _ -> raise Evarutil.NoHeadEvar + in + hrec c + let match_mode sigma m arg = match m with | ModeInput -> not (occur_existential sigma arg) @@ -543,7 +553,7 @@ struct let map_auto sigma ~secvars (k,args) concl db = let se = find k db in let st = if db.use_dn then (Some db.hintdb_state) else None in - let pat = lookup_tacs concl st se in + let pat = lookup_tacs sigma concl st se in merge_entry secvars db [] pat let map_existential sigma ~secvars (k,args) concl db = @@ -557,7 +567,7 @@ struct let se = find k db in if matches_modes sigma args se.sentry_mode then let st = if db.use_dn then Some db.hintdb_state else None in - let pat = lookup_tacs concl st se in + let pat = lookup_tacs sigma concl st se in merge_entry secvars db [] pat else merge_entry secvars db [] [] diff --git a/tactics/tactics.ml b/tactics/tactics.ml index de3572155..a29803251 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -4831,8 +4831,11 @@ let intros_transitivity n = Tacticals.New.tclTHEN intros (transitivity_gen n) (** d1 is the section variable in the global context, d2 in the goal context *) let interpretable_as_section_decl evd d1 d2 = let open Context.Named.Declaration in - let e_eq_constr_univs sigma c1 c2 = - e_eq_constr_univs sigma (EConstr.Unsafe.to_constr c1) (EConstr.Unsafe.to_constr c2) + let e_eq_constr_univs sigma c1 c2 = match eq_constr_universes !sigma c1 c2 with + | None -> false + | Some cstr -> + try ignore (Evd.add_universe_constraints !sigma cstr); true + with UniversesDiffer -> false in match d2, d1 with | LocalDef _, LocalAssum _ -> false diff --git a/toplevel/himsg.ml b/toplevel/himsg.ml index b689dfcf9..50301ee0d 100644 --- a/toplevel/himsg.ml +++ b/toplevel/himsg.ml @@ -83,8 +83,6 @@ let rec contract3' env sigma a b c = function (** Ad-hoc reductions *) -let to_unsafe_judgment j = on_judgment EConstr.Unsafe.to_constr j - let j_nf_betaiotaevar sigma j = { uj_val = Evarutil.nf_evar sigma j.uj_val; uj_type = Reductionops.nf_betaiota sigma j.uj_type } @@ -207,7 +205,7 @@ let pr_puniverses f env (c,u) = else mt()) let explain_elim_arity env sigma ind sorts c pj okinds = - let pj = to_unsafe_judgment pj in + let open EConstr in let env = make_all_name_different env in let pi = pr_inductive env (fst ind) in let pc = pr_leconstr_env env sigma c in @@ -223,7 +221,7 @@ let explain_elim_arity env sigma ind sorts c pj okinds = | WrongArity -> "wrong arity" in let ppar = pr_disjunction (fun s -> quote (pr_sort_family s)) sorts in - let ppt = pr_lconstr_env env sigma ((strip_prod_assum pj.uj_type)) in + let ppt = pr_leconstr_env env sigma (snd (decompose_prod_assum sigma pj.uj_type)) in hov 0 (str "the return type has sort" ++ spc () ++ ppt ++ spc () ++ str "while it" ++ spc () ++ str "should be " ++ ppar ++ str ".") ++ @@ -243,11 +241,10 @@ let explain_elim_arity env sigma ind sorts c pj okinds = let explain_case_not_inductive env sigma cj = let cj = Evarutil.j_nf_evar sigma cj in - let cj = to_unsafe_judgment cj in let env = make_all_name_different env in - let pc = pr_lconstr_env env sigma cj.uj_val in - let pct = pr_lconstr_env env sigma cj.uj_type in - match kind_of_term cj.uj_type with + let pc = pr_leconstr_env env sigma cj.uj_val in + let pct = pr_leconstr_env env sigma cj.uj_type in + match EConstr.kind sigma cj.uj_type with | Evar _ -> str "Cannot infer a type for this expression." | _ -> @@ -257,10 +254,9 @@ let explain_case_not_inductive env sigma cj = let explain_number_branches env sigma cj expn = let cj = Evarutil.j_nf_evar sigma cj in - let cj = to_unsafe_judgment cj in let env = make_all_name_different env in - let pc = pr_lconstr_env env sigma cj.uj_val in - let pct = pr_lconstr_env env sigma cj.uj_type in + let pc = pr_leconstr_env env sigma cj.uj_val in + let pct = pr_leconstr_env env sigma cj.uj_type in str "Matching on term" ++ brk(1,1) ++ pc ++ spc () ++ str "of type" ++ brk(1,1) ++ pct ++ spc () ++ str "expects " ++ int expn ++ str " branches." @@ -390,18 +386,16 @@ let explain_cant_apply_bad_type env sigma (n,exptyp,actualtyp) rator randl = let explain_cant_apply_not_functional env sigma rator randl = let randl = Evarutil.jv_nf_evar sigma randl in - let randl = Array.map to_unsafe_judgment randl in let rator = Evarutil.j_nf_evar sigma rator in - let rator = to_unsafe_judgment rator in let env = make_all_name_different env in let nargs = Array.length randl in (* let pe = pr_ne_context_of (str "in environment") env sigma in*) - let pr = pr_lconstr_env env sigma rator.uj_val in - let prt = pr_lconstr_env env sigma rator.uj_type in + let pr = pr_leconstr_env env sigma rator.uj_val in + let prt = pr_leconstr_env env sigma rator.uj_type in let appl = prvect_with_sep fnl (fun c -> - let pc = pr_lconstr_env env sigma c.uj_val in - let pct = pr_lconstr_env env sigma c.uj_type in + let pc = pr_leconstr_env env sigma c.uj_val in + let pct = pr_leconstr_env env sigma c.uj_type in hov 2 (pc ++ spc () ++ str ": " ++ pct)) randl in str "Illegal application (Non-functional construction): " ++ -- cgit v1.2.3 From 27fbf069ccd846383bcfb35ba1ea5bd1d95090a0 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 29 Nov 2016 23:48:28 +0100 Subject: Moving printing code from Evd to Termops. --- dev/top_printers.ml | 14 +- engine/evd.ml | 288 ----------------------------------------- engine/evd.mli | 16 --- engine/termops.ml | 324 +++++++++++++++++++++++++++++++++++++++++++++- engine/termops.mli | 36 ++++-- ltac/rewrite.ml | 2 +- pretyping/detyping.ml | 6 +- printing/printer.ml | 6 +- proofs/evar_refiner.ml | 2 +- stm/stm.ml | 2 +- toplevel/himsg.ml | 4 +- toplevel/vernacentries.ml | 4 +- 12 files changed, 370 insertions(+), 334 deletions(-) diff --git a/dev/top_printers.ml b/dev/top_printers.ml index 04aacdc09..33fc52e20 100644 --- a/dev/top_printers.ml +++ b/dev/top_printers.ml @@ -165,9 +165,9 @@ let pp_state_t n = pp (Reductionops.pr_state n) (* proof printers *) let pr_evar ev = Pp.int (Evar.repr ev) -let ppmetas metas = pp(pr_metaset metas) -let ppevm evd = pp(pr_evar_map ~with_univs:!Flags.univ_print (Some 2) evd) -let ppevmall evd = pp(pr_evar_map ~with_univs:!Flags.univ_print None evd) +let ppmetas metas = pp(Termops.pr_metaset metas) +let ppevm evd = pp(Termops.pr_evar_map ~with_univs:!Flags.univ_print (Some 2) evd) +let ppevmall evd = pp(Termops.pr_evar_map ~with_univs:!Flags.univ_print None evd) let pr_existentialset evars = prlist_with_sep spc pr_evar (Evar.Set.elements evars) let ppexistentialset evars = @@ -178,14 +178,14 @@ let ppexistentialfilter filter = match Evd.Filter.repr filter with let ppclenv clenv = pp(pr_clenv clenv) let ppgoalgoal gl = pp(Goal.pr_goal gl) let ppgoal g = pp(Printer.pr_goal g) -let ppgoalsigma g = pp(Printer.pr_goal g ++ pr_evar_map None (Refiner.project g)) +let ppgoalsigma g = pp(Printer.pr_goal g ++ Termops.pr_evar_map None (Refiner.project g)) let pphintdb db = pp(Hints.pr_hint_db db) let ppproofview p = let gls,sigma = Proofview.proofview p in - pp(pr_enum Goal.pr_goal gls ++ fnl () ++ pr_evar_map (Some 1) sigma) + pp(pr_enum Goal.pr_goal gls ++ fnl () ++ Termops.pr_evar_map (Some 1) sigma) let ppopenconstr (x : Evd.open_constr) = - let (evd,c) = x in pp (pr_evar_map (Some 2) evd ++ pr_constr c) + let (evd,c) = x in pp (Termops.pr_evar_map (Some 2) evd ++ pr_constr c) (* spiwack: deactivated until a replacement is found let pppftreestate p = pp(print_pftreestate p) *) @@ -215,7 +215,7 @@ let ppuniverse_context_set l = pp (pr_universe_context_set prlev l) let ppuniverse_subst l = pp (Univ.pr_universe_subst l) let ppuniverse_opt_subst l = pp (Universes.pr_universe_opt_subst l) let ppuniverse_level_subst l = pp (Univ.pr_universe_level_subst l) -let ppevar_universe_context l = pp (Evd.pr_evar_universe_context l) +let ppevar_universe_context l = pp (Termops.pr_evar_universe_context l) let ppconstraints c = pp (pr_constraints Level.pr c) let ppuniverseconstraints c = pp (Universes.Constraints.pr c) let ppuniverse_context_future c = diff --git a/engine/evd.ml b/engine/evd.ml index b30702d0d..3f00b91b6 100644 --- a/engine/evd.ml +++ b/engine/evd.ml @@ -258,7 +258,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 @@ -505,15 +504,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 } @@ -766,7 +756,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 = @@ -1178,280 +1167,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 (f_print_constr, print_constr_hook) = Hook.make () - -let print_constr a = protect (fun c -> Hook.get f_print_constr (Global.env ()) empty c) 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) = - 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.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_rel_decl = function - | RelDecl.LocalAssum (n,_) -> pr_name n - | RelDecl.LocalDef (n,b,_) -> str "(" ++ pr_name 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 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 () ++ - Hook.get f_print_constr env empty t1 ++ spc () ++ - str (match pbty with - | Reduction.CONV -> "==" - | Reduction.CUMUL -> "<=") ++ - spc () ++ Hook.get f_print_constr env empty 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 eeb9fd861..fc1d4a514 100644 --- a/engine/evd.mli +++ b/engine/evd.mli @@ -584,22 +584,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 print_constr_hook : (Environ.env -> evar_map -> constr -> Pp.std_ppcmds) Hook.t -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/termops.ml b/engine/termops.ml index c35e92f97..ff1a0d9de 100644 --- a/engine/termops.ml +++ b/engine/termops.ml @@ -101,7 +101,329 @@ let term_printer = ref (fun _env _sigma c -> pr_constr (EConstr.Unsafe.to_constr 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 -let () = Hook.set Evd.print_constr_hook (fun env sigma c -> !term_printer env sigma (EConstr.of_constr c)) + +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.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 (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.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,_) -> pr_name n + | RelDecl.LocalDef (n,b,_) -> str "(" ++ pr_name 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 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 Evd.empty (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 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 diff --git a/engine/termops.mli b/engine/termops.mli index 841de34b1..512bb05ff 100644 --- a/engine/termops.mli +++ b/engine/termops.mli @@ -20,15 +20,6 @@ val print_sort : sorts -> std_ppcmds val pr_sort_family : sorts_family -> std_ppcmds val pr_fix : ('a -> std_ppcmds) -> ('a, 'a) pfixpoint -> std_ppcmds -(** 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 - (** about contexts *) val push_rel_assum : Name.t * types -> env -> env val push_rels_assum : (Name.t * Constr.types) list -> env -> env @@ -287,3 +278,30 @@ val on_judgment_type : ('t -> 't) -> ('c, 't) punsafe_judgment -> ('c, 't) puns (** {6 Functions to deal with impossible cases } *) val set_impossible_default_clause : (unit -> (Constr.constr * Constr.types) Univ.in_universe_context_set) -> unit val coq_unit_judge : unit -> unsafe_judgment Univ.in_universe_context_set + +(** {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_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 + +(** 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/ltac/rewrite.ml b/ltac/rewrite.ml index b45ead217..bc57fcf56 100644 --- a/ltac/rewrite.ml +++ b/ltac/rewrite.ml @@ -1511,7 +1511,7 @@ let cl_rewrite_clause_aux ?(abs=None) strat env avoid sigma concl is_hyp : resul if not (Evd.is_defined acc ev) then user_err ~hdr:"rewrite" (str "Unsolved constraint remaining: " ++ spc () ++ - Evd.pr_evar_info (Evd.find acc ev)) + Termops.pr_evar_info (Evd.find acc ev)) else Evd.remove acc ev) cstrs evars' in diff --git a/pretyping/detyping.ml b/pretyping/detyping.ml index d4e156fa4..c8d945c0b 100644 --- a/pretyping/detyping.ml +++ b/pretyping/detyping.ml @@ -422,7 +422,7 @@ let detype_sort sigma = function | Type u -> GType (if !print_universes - then [dl, Pp.string_of_ppcmds (Univ.Universe.pr_with (Evd.pr_evd_level sigma) u)] + then [dl, Pp.string_of_ppcmds (Univ.Universe.pr_with (Termops.pr_evd_level sigma) u)] else []) type binder_kind = BProd | BLambda | BLetIn @@ -434,7 +434,7 @@ let detype_anonymous = ref (fun loc n -> anomaly ~label:"detype" (Pp.str "index let set_detype_anonymous f = detype_anonymous := f let detype_level sigma l = - GType (Some (dl, Pp.string_of_ppcmds (Evd.pr_evd_level sigma l))) + GType (Some (dl, Pp.string_of_ppcmds (Termops.pr_evd_level sigma l))) let detype_instance sigma l = if Univ.Instance.is_empty l then None @@ -533,7 +533,7 @@ let rec detype flags avoid env sigma t = let id,l = try let id = match Evd.evar_ident evk sigma with - | None -> Evd.pr_evar_suggested_name evk sigma + | None -> Termops.pr_evar_suggested_name evk sigma | Some id -> id in let l = Evd.evar_instance_array bound_to_itself_or_letin (Evd.find sigma evk) cl in diff --git a/printing/printer.ml b/printing/printer.ml index d9060c172..c7d3a1ba3 100644 --- a/printing/printer.ml +++ b/printing/printer.ml @@ -223,7 +223,7 @@ let safe_pr_constr t = let pr_universe_ctx sigma c = if !Detyping.print_universes && not (Univ.UContext.is_empty c) then fnl()++pr_in_comment (fun c -> v 0 - (Univ.pr_universe_context (Evd.pr_evd_level sigma) c)) c + (Univ.pr_universe_context (Termops.pr_evd_level sigma) c)) c else mt() @@ -240,7 +240,7 @@ let pr_puniverses f env (c,u) = else mt ()) let pr_constant env cst = pr_global_env (Termops.vars_of_env env) (ConstRef cst) -let pr_existential_key = Evd.pr_existential_key +let pr_existential_key = Termops.pr_existential_key let pr_existential env sigma ev = pr_lconstr_env env sigma (mkEvar ev) let pr_inductive env ind = pr_lconstr_env env Evd.empty (mkInd ind) let pr_constructor env cstr = pr_lconstr_env env Evd.empty (mkConstruct cstr) @@ -922,4 +922,4 @@ let pr_polymorphic b = let pr_universe_instance evd ctx = let inst = Univ.UContext.instance ctx in - str"@{" ++ Univ.Instance.pr (Evd.pr_evd_level evd) inst ++ str"}" + str"@{" ++ Univ.Instance.pr (Termops.pr_evd_level evd) inst ++ str"}" diff --git a/proofs/evar_refiner.ml b/proofs/evar_refiner.ml index 1ddb67edd..59a41792a 100644 --- a/proofs/evar_refiner.ml +++ b/proofs/evar_refiner.ml @@ -58,6 +58,6 @@ let w_refine (evk,evi) (ltac_var,rawc) sigma = let loc = Glob_ops.loc_of_glob_constr rawc in user_err ~loc (str "Instance is not well-typed in the environment of " ++ - pr_existential_key sigma evk ++ str ".") + Termops.pr_existential_key sigma evk ++ str ".") in define_and_solve_constraints evk typed_c env (evars_reset_evd sigma' sigma) diff --git a/stm/stm.ml b/stm/stm.ml index 47e806929..6f5da37f4 100644 --- a/stm/stm.ml +++ b/stm/stm.ml @@ -1791,7 +1791,7 @@ end = struct (* {{{ *) prerr_endline (fun () -> string_of_ppcmds(hov 0 ( str"g=" ++ int (Evar.repr gid) ++ spc () ++ str"t=" ++ (Printer.pr_constr pt) ++ spc () ++ - str"uc=" ++ Evd.pr_evar_universe_context uc))); + str"uc=" ++ Termops.pr_evar_universe_context uc))); (if abstract then Tactics.tclABSTRACT None else (fun x -> x)) (V82.tactic (Refiner.tclPUSHEVARUNIVCONTEXT uc) <*> Tactics.exact_no_check (EConstr.of_constr pt)) diff --git a/toplevel/himsg.ml b/toplevel/himsg.ml index 50301ee0d..6982205fd 100644 --- a/toplevel/himsg.ml +++ b/toplevel/himsg.ml @@ -691,7 +691,7 @@ let explain_non_linear_unification env sigma m t = let explain_unsatisfied_constraints env sigma cst = strbrk "Unsatisfied constraints: " ++ - Univ.pr_constraints (Evd.pr_evd_level sigma) cst ++ + Univ.pr_constraints (Termops.pr_evd_level sigma) cst ++ spc () ++ str "(maybe a bugged tactic)." let explain_type_error env sigma err = @@ -903,7 +903,7 @@ let explain_not_match_error = function quote (Printer.safe_pr_lconstr_env env Evd.empty t2) | IncompatibleConstraints cst -> str " the expected (polymorphic) constraints do not imply " ++ - quote (Univ.pr_constraints (Evd.pr_evd_level Evd.empty) cst) + quote (Univ.pr_constraints (Termops.pr_evd_level Evd.empty) cst) let explain_signature_mismatch l spec why = str "Signature components for label " ++ pr_label l ++ diff --git a/toplevel/vernacentries.ml b/toplevel/vernacentries.ml index f3d73b76c..55477cbcf 100644 --- a/toplevel/vernacentries.ml +++ b/toplevel/vernacentries.ml @@ -81,8 +81,8 @@ let show_universes () = let gls = Proof.V82.subgoals pfts in let sigma = gls.Evd.sigma in let ctx = Evd.universe_context_set (Evd.nf_constraints sigma) in - Feedback.msg_notice (Evd.pr_evar_universe_context (Evd.evar_universe_context sigma)); - Feedback.msg_notice (str"Normalized constraints: " ++ Univ.pr_universe_context_set (Evd.pr_evd_level sigma) ctx) + Feedback.msg_notice (Termops.pr_evar_universe_context (Evd.evar_universe_context sigma)); + Feedback.msg_notice (str"Normalized constraints: " ++ Univ.pr_universe_context_set (Termops.pr_evd_level sigma) ctx) let show_prooftree () = (* Spiwack: proof tree is currently not working *) -- cgit v1.2.3 From 1683b718f85134fdb0d49535e489344e1a7d56f5 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Wed, 30 Nov 2016 00:33:08 +0100 Subject: Making Evd independent from Namegen. --- engine/evarutil.ml | 10 +++++++++- engine/evd.ml | 27 ++++++++++----------------- engine/evd.mli | 2 +- engine/sigma.ml | 4 ++-- engine/sigma.mli | 2 +- 5 files changed, 23 insertions(+), 22 deletions(-) diff --git a/engine/evarutil.ml b/engine/evarutil.ml index 0582e34be..e94fd72f1 100644 --- a/engine/evarutil.ml +++ b/engine/evarutil.ml @@ -418,6 +418,14 @@ let new_pure_evar sign evd ?(src=default_source) ?(filter = Filter.identity) ?ca 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; @@ -427,7 +435,7 @@ 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 diff --git a/engine/evd.ml b/engine/evd.ml index 3f00b91b6..b8f4b126a 100644 --- a/engine/evd.ml +++ b/engine/evd.ml @@ -365,7 +365,7 @@ 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 @@ -379,20 +379,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 (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 @@ -462,9 +455,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 +474,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 = diff --git a/engine/evd.mli b/engine/evd.mli index fc1d4a514..2151000b6 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 diff --git a/engine/sigma.ml b/engine/sigma.ml index 9381a33af..001f8be80 100644 --- a/engine/sigma.ml +++ b/engine/sigma.ml @@ -29,8 +29,8 @@ let here x s = Sigma (x, s, ()) 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 +let new_evar sigma ?name info = + let (sigma, evk) = Evd.new_evar sigma ?name info in Fresh (evk, sigma, ()) let define evk c sigma = diff --git a/engine/sigma.mli b/engine/sigma.mli index 7473a251b..8e8df02f2 100644 --- a/engine/sigma.mli +++ b/engine/sigma.mli @@ -61,7 +61,7 @@ val to_evar : 'r evar -> Evar.t type 'r fresh = Fresh : 's evar * 's t * ('r, 's) le -> 'r fresh -val new_evar : 'r t -> ?naming:Misctypes.intro_pattern_naming_expr -> +val new_evar : 'r t -> ?name:Id.t -> Evd.evar_info -> 'r fresh val define : 'r evar -> Constr.t -> 'r t -> (unit, 'r) sigma -- cgit v1.2.3 From be51c33a6bf91a00fdd5f3638ddb5b3cc3a2c626 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Wed, 30 Nov 2016 00:41:31 +0100 Subject: Namegen primitives now apply on evar constrs. Incidentally, this fixes a printing bug in output/inference.v where the displayed name of an evar was the wrong one because its type was not evar-expanded enough. --- engine/eConstr.ml | 2 + engine/eConstr.mli | 1 + engine/engine.mllib | 2 +- engine/evarutil.ml | 10 +-- engine/evarutil.mli | 4 +- engine/namegen.ml | 116 +++++++++++++++++--------------- engine/namegen.mli | 44 ++++++------ engine/termops.ml | 10 +-- engine/termops.mli | 2 +- interp/constrextern.ml | 8 ++- interp/constrintern.ml | 2 +- interp/notation_ops.ml | 1 + interp/reserve.ml | 2 + library/impargs.ml | 8 ++- ltac/evar_tactics.ml | 6 +- ltac/extratactics.ml4 | 2 - plugins/decl_mode/decl_interp.ml | 2 +- plugins/decl_mode/decl_proof_instr.ml | 18 ++--- plugins/funind/glob_term_to_relation.ml | 9 ++- plugins/funind/merge.ml | 2 + plugins/funind/recdef.ml | 4 +- pretyping/cases.ml | 15 ++--- pretyping/detyping.ml | 104 ++++++++++++++-------------- pretyping/detyping.mli | 7 +- pretyping/evardefine.ml | 2 +- pretyping/indrec.ml | 20 +++++- pretyping/inductiveops.ml | 32 +++++---- pretyping/inductiveops.mli | 6 +- pretyping/pretyping.ml | 75 ++++++++++----------- pretyping/tacred.ml | 2 +- pretyping/typing.ml | 3 +- pretyping/unification.ml | 5 +- printing/printer.ml | 2 +- proofs/clenv.ml | 6 +- tactics/class_tactics.ml | 2 +- tactics/eqschemes.ml | 16 ++++- tactics/equality.ml | 17 ++--- tactics/inv.ml | 9 +-- tactics/leminv.ml | 7 +- tactics/tactics.ml | 27 ++++---- test-suite/output/inference.out | 12 ++-- toplevel/himsg.ml | 42 ++++++------ toplevel/vernacentries.ml | 2 +- 43 files changed, 351 insertions(+), 317 deletions(-) diff --git a/engine/eConstr.ml b/engine/eConstr.ml index 657285de2..7d4d17c63 100644 --- a/engine/eConstr.ml +++ b/engine/eConstr.ml @@ -131,6 +131,8 @@ 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 diff --git a/engine/eConstr.mli b/engine/eConstr.mli index d6b2113d2..83536d6f8 100644 --- a/engine/eConstr.mli +++ b/engine/eConstr.mli @@ -114,6 +114,7 @@ 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 diff --git a/engine/engine.mllib b/engine/engine.mllib index c78f37a85..1b670d366 100644 --- a/engine/engine.mllib +++ b/engine/engine.mllib @@ -1,10 +1,10 @@ Logic_monad -Namegen Universes UState Evd Sigma EConstr +Namegen Termops Proofview_monad Evarutil diff --git a/engine/evarutil.ml b/engine/evarutil.ml index e94fd72f1..ce6d06f23 100644 --- a/engine/evarutil.ml +++ b/engine/evarutil.ml @@ -326,7 +326,7 @@ let push_var id (n, s) = 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 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 = @@ -354,7 +354,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 (EConstr.Unsafe.to_constr (RelDecl.get_type decl)) 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) -> @@ -375,7 +375,7 @@ let push_rel_decl_to_named_context decl (subst, vsubst, avoid, nc) = 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 @@ -390,7 +390,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) @@ -452,7 +452,7 @@ let new_evar_instance sign evd typ ?src ?filter ?candidates ?store ?naming ?prin (* [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 sign,typ',instance,subst,vsubst = push_rel_context_to_named_context env (Sigma.to_evar_map 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 = diff --git a/engine/evarutil.mli b/engine/evarutil.mli index 2da4f8043..da49d4e11 100644 --- a/engine/evarutil.mli +++ b/engine/evarutil.mli @@ -216,9 +216,9 @@ type ext_named_context = Id.Set.t * named_context val push_rel_decl_to_named_context : - rel_declaration -> 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/namegen.ml b/engine/namegen.ml index e56ec2877..7b7302957 100644 --- a/engine/namegen.ml +++ b/engine/namegen.ml @@ -15,12 +15,13 @@ open Util open Names open Term +open Environ +open EConstr open Vars open Nametab open Nameops open Libnames open Globnames -open Environ open Context.Rel.Declaration module RelDecl = Context.Rel.Declaration @@ -84,13 +85,20 @@ let is_section_variable id = (**********************************************************************) (* 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) @@ -105,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))) @@ -119,7 +127,7 @@ let hdchar env c = | Rel n -> (if n<=k then "p" (* the initial term is flexible product/function *) else - try match Environ.lookup_rel (n-k) env with + 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") @@ -131,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 *) @@ -185,10 +193,10 @@ let restart_subscript id = *** make_ident id (Some 0) *** but compatibility would be lost... *) forget_subscript id -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 @@ -218,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 @@ -228,9 +236,9 @@ 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 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 @@ -292,7 +300,7 @@ 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 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 @@ -300,7 +308,7 @@ let make_all_name_different env = let env0 = reset_with_named_context sign env in Context.Rel.fold_outside (fun decl newenv -> - let na = named_hd newenv (RelDecl.get_type decl) (RelDecl.get_name decl) 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 (RelDecl.set_name (Name id) decl) newenv) @@ -311,12 +319,12 @@ let make_all_name_different env = 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 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 -> @@ -324,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 *) @@ -349,45 +357,45 @@ 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' (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' (na' :: env) c2) | Cast (c,k,t) -> mkCast (rename avoid env c, k,t) diff --git a/engine/namegen.mli b/engine/namegen.mli index 33ce9a34d..058b1c228 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 *) @@ -98,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/termops.ml b/engine/termops.ml index ff1a0d9de..d8e712abc 100644 --- a/engine/termops.ml +++ b/engine/termops.ml @@ -115,7 +115,7 @@ let pr_evar_suggested_name evk sigma = | _,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 + 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 @@ -316,7 +316,7 @@ let print_env_short env = str "[" ++ pr_sequence pr_named_decl nc ++ str "]" ++ spc () ++ str "[" ++ pr_sequence pr_rel_decl rc ++ str "]" -let pr_evar_constraints pbs = +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 @@ -326,10 +326,10 @@ let pr_evar_constraints pbs = 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 + Namegen.make_all_name_different env sigma in print_env_short env ++ spc () ++ str "|-" ++ spc () ++ - print_constr_env env Evd.empty (EConstr.of_constr t1) ++ spc () ++ + print_constr_env env sigma (EConstr.of_constr t1) ++ spc () ++ str (match pbty with | Reduction.CONV -> "==" | Reduction.CUMUL -> "<=") ++ @@ -346,7 +346,7 @@ let pr_evar_map_gen with_univs pr_evars sigma = if List.is_empty conv_pbs then mt () else str "CONSTRAINTS:" ++ brk (0, 1) ++ - pr_evar_constraints conv_pbs ++ fnl () + pr_evar_constraints sigma conv_pbs ++ fnl () and metas = if List.is_empty (Evd.meta_list sigma) then mt () else diff --git a/engine/termops.mli b/engine/termops.mli index 512bb05ff..0dd5d96fb 100644 --- a/engine/termops.mli +++ b/engine/termops.mli @@ -288,7 +288,7 @@ 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_constraint list -> 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 diff --git a/interp/constrextern.ml b/interp/constrextern.ml index 3077231be..8debc06bb 100644 --- a/interp/constrextern.ml +++ b/interp/constrextern.ml @@ -953,6 +953,7 @@ let extern_constr_gen lax goal_concl_style scopt env sigma t = (* Not "goal_concl_style" means do alpha-conversion avoiding only *) (* those goal/section/rel variables that occurs in the subterm under *) (* consideration; see namegen.ml for further details *) + let t = EConstr.of_constr t in let avoid = if goal_concl_style then ids_of_context env else [] in let r = Detyping.detype ~lax:lax goal_concl_style avoid env sigma t in let vars = vars_of_env env in @@ -965,6 +966,7 @@ let extern_constr ?(lax=false) goal_concl_style env sigma t = extern_constr_gen lax goal_concl_style None env sigma t let extern_type goal_concl_style env sigma t = + let t = EConstr.of_constr t in let avoid = if goal_concl_style then ids_of_context env else [] in let r = Detyping.detype goal_concl_style avoid env sigma t in extern_glob_type (vars_of_env env) r @@ -1042,14 +1044,16 @@ let rec glob_of_pat env sigma = function | _ -> anomaly (Pp.str "PCase with non-trivial predicate but unknown inductive") in GCases (loc,RegularStyle,rtn,[glob_of_pat env sigma tm,indnames],mat) - | PFix f -> Detyping.detype_names false [] env (Global.env()) sigma (mkFix f) (** FIXME bad env *) - | PCoFix c -> Detyping.detype_names false [] env (Global.env()) sigma (mkCoFix c) + | PFix f -> Detyping.detype_names false [] env (Global.env()) sigma (EConstr.of_constr (mkFix f)) (** FIXME bad env *) + | PCoFix c -> Detyping.detype_names false [] env (Global.env()) sigma (EConstr.of_constr (mkCoFix c)) | PSort s -> GSort (loc,s) let extern_constr_pattern env sigma pat = extern true (None,[]) Id.Set.empty (glob_of_pat env sigma pat) let extern_rel_context where env sigma sign = + let sign = List.map (fun d -> Termops.map_rel_decl EConstr.of_constr d) sign in + let where = Option.map EConstr.of_constr where in let a = detype_rel_context where [] (names_of_rel_context env,env) sigma sign in let vars = vars_of_env env in let a = List.map (fun (p,bk,x,t) -> (Inl p,bk,x,t)) a in diff --git a/interp/constrintern.ml b/interp/constrintern.ml index 41e2e4e6f..bac97aa3f 100644 --- a/interp/constrintern.ml +++ b/interp/constrintern.ml @@ -1843,7 +1843,7 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c = (add_name match_acc (loc,x)) ((loc,x)::var_acc) | (LocalAssum (cano_name,ty) | LocalDef (cano_name,_,ty)) :: t, c::tt -> let fresh = - Namegen.next_name_away_with_default_using_types "iV" cano_name forbidden_names ty in + Namegen.next_name_away_with_default_using_types "iV" cano_name forbidden_names (EConstr.of_constr ty) in canonize_args t tt (fresh::forbidden_names) ((fresh,c)::match_acc) ((cases_pattern_loc c,Name fresh)::var_acc) | _ -> assert false in diff --git a/interp/notation_ops.ml b/interp/notation_ops.ml index 7dbd94aa7..549e8e787 100644 --- a/interp/notation_ops.ml +++ b/interp/notation_ops.ml @@ -441,6 +441,7 @@ let notation_constr_of_glob_constr nenv a = (* Substitution of kernel names, avoiding a list of bound identifiers *) let notation_constr_of_constr avoiding t = + let t = EConstr.of_constr t in let t = Detyping.detype false avoiding (Global.env()) Evd.empty t in let nenv = { ninterp_var_type = Id.Map.empty; diff --git a/interp/reserve.ml b/interp/reserve.ml index a4d4f4027..1565ba4a9 100644 --- a/interp/reserve.ml +++ b/interp/reserve.ml @@ -107,7 +107,9 @@ let constr_key c = let revert_reserved_type t = try + let t = EConstr.Unsafe.to_constr t in let reserved = KeyMap.find (constr_key t) !reserve_revtable in + let t = EConstr.of_constr t in let t = Detyping.detype false [] (Global.env()) Evd.empty t in (* pedrot: if [Notation_ops.match_notation_constr] may raise [Failure _] then I've introduced a bug... *) diff --git a/library/impargs.ml b/library/impargs.ml index 836568b89..a63264b66 100644 --- a/library/impargs.ml +++ b/library/impargs.ml @@ -228,10 +228,12 @@ let rec is_rigid_head t = match kind_of_term t with (* calcule la liste des arguments implicites *) -let find_displayed_name_in all avoid na (_,b as envnames_b) = +let find_displayed_name_in all avoid na (env, b) = + let b = EConstr.of_constr b in + let envnames_b = (env, b) in let flag = RenamingElsewhereFor envnames_b in - if all then compute_and_force_displayed_name_in flag avoid na b - else compute_displayed_name_in flag avoid na b + if all then compute_and_force_displayed_name_in Evd.empty flag avoid na b + else compute_displayed_name_in Evd.empty flag avoid na b let compute_implicits_gen strict strongly_strict revpat contextual all env t = let rigid = ref true in diff --git a/ltac/evar_tactics.ml b/ltac/evar_tactics.ml index 01cff94a8..5d3f6df03 100644 --- a/ltac/evar_tactics.ml +++ b/ltac/evar_tactics.ml @@ -88,14 +88,14 @@ let let_evar name typ = let env = Proofview.Goal.env gl in let sigma = ref sigma in let _ = Typing.e_sort_of env sigma typ in - let sigma = Sigma.Unsafe.of_evar_map !sigma in + let sigma = !sigma in let id = match name with | Names.Anonymous -> - let typ = EConstr.Unsafe.to_constr typ in - let id = Namegen.id_of_name_using_hdchar env typ name in + let id = Namegen.id_of_name_using_hdchar env sigma typ name in Namegen.next_ident_away_in_goal id (Termops.ids_of_named_context (Environ.named_context env)) | Names.Name id -> id in + let sigma = Sigma.Unsafe.of_evar_map sigma in let Sigma (evar, sigma, p) = Evarutil.new_evar env sigma ~src ~naming:(Misctypes.IntroFresh id) typ in let tac = (Tactics.letin_tac None (Names.Name id) evar None Locusops.nowhere) diff --git a/ltac/extratactics.ml4 b/ltac/extratactics.ml4 index 34159d945..58c5823c5 100644 --- a/ltac/extratactics.ml4 +++ b/ltac/extratactics.ml4 @@ -645,8 +645,6 @@ let subst_hole_with_term occ tc t = open Tacmach let hResolve id c occ t = - let c = EConstr.Unsafe.to_constr c in - let t = EConstr.Unsafe.to_constr t in Proofview.Goal.nf_s_enter { s_enter = begin fun gl -> let sigma = Proofview.Goal.sigma gl in let sigma = Sigma.to_evar_map sigma in diff --git a/plugins/decl_mode/decl_interp.ml b/plugins/decl_mode/decl_interp.ml index ae057b458..828d634d0 100644 --- a/plugins/decl_mode/decl_interp.ml +++ b/plugins/decl_mode/decl_interp.ml @@ -316,7 +316,7 @@ let rec match_aliases names constr = function let args,bnames,body = match_aliases qnames body q in st::args,bnames,body -let detype_ground env c = Detyping.detype false [] env Evd.empty c +let detype_ground env c = Detyping.detype false [] env Evd.empty (EConstr.of_constr c) let interp_cases info env sigma params (pat:cases_pattern_expr) hyps = let et,pinfo = diff --git a/plugins/decl_mode/decl_proof_instr.ml b/plugins/decl_mode/decl_proof_instr.ml index adc4ad8a3..4fe59d755 100644 --- a/plugins/decl_mode/decl_proof_instr.ml +++ b/plugins/decl_mode/decl_proof_instr.ml @@ -440,7 +440,7 @@ let concl_refiner metas body gls = [] -> anomaly ~label:"concl_refiner" (Pp.str "cannot happen") | (n,typ)::rest -> let _A = subst_meta subst typ in - let x = id_of_name_using_hdchar env _A Anonymous in + let x = id_of_name_using_hdchar env evd (EConstr.of_constr _A) Anonymous in let _x = fresh_id avoid x gls in let nenv = Environ.push_named (LocalAssum (_x,_A)) env in let asort = family_of_sort (Typing.e_sort_of nenv (ref evd) (EConstr.of_constr _A)) in @@ -895,7 +895,6 @@ let build_per_info etype casee gls = let ctyp=pf_unsafe_type_of gls (EConstr.of_constr casee) in let is_dep = dependent (project gls) (EConstr.of_constr casee) concl in let hd,args = decompose_app (special_whd gls ctyp) in - let ctyp = EConstr.Unsafe.to_constr ctyp in let (ind,u) = try destInd hd @@ -909,10 +908,11 @@ let build_per_info etype casee gls = let params,real_args = List.chop nparams args in let abstract_obj c body = let typ=pf_unsafe_type_of gls (EConstr.of_constr c) in - let typ = EConstr.Unsafe.to_constr typ in - lambda_create env (typ,subst_term (project gls) (EConstr.of_constr c) (EConstr.of_constr body)) in + lambda_create env (project gls) (typ,Termops.subst_term (project gls) (EConstr.of_constr c) body) in let pred= List.fold_right abstract_obj - real_args (lambda_create env (ctyp,subst_term (project gls) (EConstr.of_constr casee) concl)) in + real_args (lambda_create env (project gls) (ctyp,Termops.subst_term (project gls) (EConstr.of_constr casee) concl)) in + let ctyp = EConstr.Unsafe.to_constr ctyp in + let pred = EConstr.Unsafe.to_constr pred in is_dep, {per_casee=casee; per_ctype=ctyp; @@ -1275,16 +1275,15 @@ let rec execute_cases fix_name per_info tacnext args objs nhrec tree gls = let env=pf_env gls in let ctyp=pf_unsafe_type_of gls (EConstr.of_constr casee) in let hd,all_args = decompose_app (special_whd gls ctyp) in - let ctyp = EConstr.Unsafe.to_constr ctyp in let ind', u = destInd hd in let _ = assert (eq_ind ind' ind) in (* just in case *) let params,real_args = List.chop nparams all_args in let abstract_obj c body = let typ=pf_unsafe_type_of gls (EConstr.of_constr c) in - let typ = EConstr.Unsafe.to_constr typ in - lambda_create env (typ,subst_term (project gls) (EConstr.of_constr c) (EConstr.of_constr body)) in + lambda_create env (project gls) (typ,Termops.subst_term (project gls) (EConstr.of_constr c) body) in let elim_pred = List.fold_right abstract_obj - real_args (lambda_create env (ctyp,subst_term (project gls) (EConstr.of_constr casee) concl)) in + real_args (lambda_create env (project gls) (ctyp,Termops.subst_term (project gls) (EConstr.of_constr casee) concl)) in + let elim_pred = EConstr.Unsafe.to_constr elim_pred in let case_info = Inductiveops.make_case_info env ind RegularStyle in let gen_arities = Inductive.arities_of_constructors (ind,u) spec in let f_ids typ = @@ -1345,6 +1344,7 @@ let rec execute_cases fix_name per_info tacnext args objs nhrec tree gls = let understand_my_constr env sigma c concl = let env = env in + let c = EConstr.of_constr c in let rawc = Detyping.detype false [] env Evd.empty c in let rec frob = function | GEvar _ -> GHole (Loc.ghost,Evar_kinds.QuestionMark Evar_kinds.Expand,Misctypes.IntroAnonymous,None) diff --git a/plugins/funind/glob_term_to_relation.ml b/plugins/funind/glob_term_to_relation.ml index fd2f4bbd3..63bd3848f 100644 --- a/plugins/funind/glob_term_to_relation.ml +++ b/plugins/funind/glob_term_to_relation.ml @@ -421,7 +421,7 @@ let rec pattern_to_term_and_type env typ = function Array.to_list (Array.init (cst_narg - List.length patternl) - (fun i -> Detyping.detype false [] env (Evd.from_env env) csta.(i)) + (fun i -> Detyping.detype false [] env (Evd.from_env env) (EConstr.of_constr csta.(i))) ) in let patl_as_term = @@ -504,7 +504,6 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return = *) let rt_as_constr,ctx = Pretyping.understand env (Evd.from_env env) rt in let rt_typ = Typing.unsafe_type_of env (Evd.from_env env) (EConstr.of_constr rt_as_constr) in - let rt_typ = EConstr.Unsafe.to_constr rt_typ in let res_raw_type = Detyping.detype false [] env (Evd.from_env env) rt_typ in let res = fresh_id args_res.to_avoid "_res" in let new_avoid = res::args_res.to_avoid in @@ -757,7 +756,6 @@ and build_entry_lc_from_case_term env types funname make_discr patterns_to_preve let typ_of_id = Typing.unsafe_type_of env_with_pat_ids (Evd.from_env env) (EConstr.mkVar id) in - let typ_of_id = EConstr.Unsafe.to_constr typ_of_id in let raw_typ_of_id = Detyping.detype false [] env_with_pat_ids (Evd.from_env env) typ_of_id @@ -804,6 +802,7 @@ and build_entry_lc_from_case_term env types funname make_discr patterns_to_preve List.map3 (fun pat e typ_as_constr -> let this_pat_ids = ids_of_pat pat in + let typ_as_constr = EConstr.of_constr typ_as_constr in let typ = Detyping.detype false [] new_env (Evd.from_env env) typ_as_constr in let pat_as_term = pattern_to_term pat in List.fold_right @@ -811,7 +810,6 @@ and build_entry_lc_from_case_term env types funname make_discr patterns_to_preve if Id.Set.mem id this_pat_ids then (Prod (Name id), let typ_of_id = Typing.unsafe_type_of new_env (Evd.from_env env) (EConstr.mkVar id) in - let typ_of_id = EConstr.Unsafe.to_constr typ_of_id in let raw_typ_of_id = Detyping.detype false [] new_env (Evd.from_env env) typ_of_id in @@ -970,7 +968,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt = (List.map (fun p -> Detyping.detype false [] env (Evd.from_env env) - p) params)@(Array.to_list + (EConstr.of_constr p)) params)@(Array.to_list (Array.make (List.length args' - nparam) (mkGHole ())))) @@ -988,6 +986,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt = let ty' = snd (Util.List.chop nparam ty) in List.fold_left2 (fun acc var_as_constr arg -> + let arg = EConstr.of_constr arg in if isRel var_as_constr then let na = RelDecl.get_name (Environ.lookup_rel (destRel var_as_constr) env) in diff --git a/plugins/funind/merge.ml b/plugins/funind/merge.ml index 691385fad..44aacaf45 100644 --- a/plugins/funind/merge.ml +++ b/plugins/funind/merge.ml @@ -777,6 +777,7 @@ let merge_inductive_body (shift:merge_infos) avoid (oib1:one_inductive_body) let mkrawcor nme avoid typ = (* first replace rel 1 by a varname *) let substindtyp = substitterm 0 (mkRel 1) (mkVar nme) typ in + let substindtyp = EConstr.of_constr substindtyp in Detyping.detype false (Id.Set.elements avoid) (Global.env()) Evd.empty substindtyp in let lcstr1: glob_constr list = Array.to_list (Array.map (mkrawcor ind1name avoid) oib1.mind_user_lc) in @@ -860,6 +861,7 @@ let glob_constr_list_to_inductive_expr prms1 prms2 mib1 mib2 shift let mkProd_reldecl (rdecl:Context.Rel.Declaration.t) (t2:glob_constr) = match rdecl with | LocalAssum (nme,t) -> + let t = EConstr.of_constr t in let traw = Detyping.detype false [] (Global.env()) Evd.empty t in GProd (Loc.ghost,nme,Explicit,traw,t2) | LocalDef _ -> assert false diff --git a/plugins/funind/recdef.ml b/plugins/funind/recdef.ml index f5ea32878..2a66ba852 100644 --- a/plugins/funind/recdef.ml +++ b/plugins/funind/recdef.ml @@ -122,8 +122,8 @@ let pf_get_new_ids idl g = [] let compute_renamed_type gls c = - EConstr.of_constr (rename_bound_vars_as_displayed (*no avoid*) [] (*no rels*) [] - (EConstr.Unsafe.to_constr (pf_unsafe_type_of gls c))) + rename_bound_vars_as_displayed (project gls) (*no avoid*) [] (*no rels*) [] + (pf_unsafe_type_of gls c) let h'_id = Id.of_string "h'" let teq_id = Id.of_string "teq" let ano_id = Id.of_string "anonymous" diff --git a/pretyping/cases.ml b/pretyping/cases.ml index a5a5fe6d2..490bc2801 100644 --- a/pretyping/cases.ml +++ b/pretyping/cases.ml @@ -708,7 +708,7 @@ let merge_name get_name obj = function let merge_names get_name = List.map2 (merge_name get_name) -let get_names env sign eqns = +let get_names env sigma sign eqns = let names1 = List.make (Context.Rel.length sign) Anonymous in (* If any, we prefer names used in pats, from top to bottom *) let names2,aliasname = @@ -727,7 +727,7 @@ let get_names env sign eqns = (fun (l,avoid) d na -> let na = merge_name - (fun (LocalAssum (na,t) | LocalDef (na,_,t)) -> Name (next_name_away (named_hd env (EConstr.Unsafe.to_constr t) na) avoid)) + (fun (LocalAssum (na,t) | LocalDef (na,_,t)) -> Name (next_name_away (named_hd env sigma t na) avoid)) d na in (na::l,(out_name na)::avoid)) @@ -924,7 +924,7 @@ let rec extract_predicate ccl = function ccl let abstract_predicate env sigma indf cur realargs (names,na) tms ccl = - let sign = make_arity_signature env true indf in + let sign = make_arity_signature env sigma true indf in (* n is the number of real args + 1 (+ possible let-ins in sign) *) let n = List.length sign in (* Before abstracting we generalize over cur and on those realargs *) @@ -945,7 +945,7 @@ let abstract_predicate env sigma indf cur realargs (names,na) tms ccl = let pred = extract_predicate ccl tms in (* Build the predicate properly speaking *) let sign = List.map2 set_name (na::names) sign in - EConstr.of_constr (it_mkLambda_or_LetIn_name env (EConstr.Unsafe.to_constr pred) sign) + it_mkLambda_or_LetIn_name env sigma pred sign (* [expand_arg] is used by [specialize_predicate] if Yk denotes [Xk;xk] or [Xk], @@ -1238,7 +1238,7 @@ let build_branch initial current realargs deps (realnames,curname) pb arsign eqn (* that had matched constructor C *) let cs_args = const_info.cs_args in let cs_args = List.map (fun d -> map_rel_decl EConstr.of_constr d) cs_args in - let names,aliasname = get_names pb.env cs_args eqns in + let names,aliasname = get_names pb.env !(pb.evdref) cs_args eqns in let typs = List.map2 RelDecl.set_name names cs_args in @@ -1714,7 +1714,7 @@ let build_tycon loc env tycon_env s subst tycon extenv evdref t = let build_inversion_problem loc env sigma tms t = let make_patvar t (subst,avoid) = - let id = next_name_away (named_hd env (EConstr.Unsafe.to_constr t) Anonymous) avoid in + let id = next_name_away (named_hd env sigma t Anonymous) avoid in PatVar (Loc.ghost,Name id), ((id,t)::subst, id::avoid) in let rec reveal_pattern t (subst,avoid as acc) = match EConstr.kind sigma (whd_all env sigma t) with @@ -1733,8 +1733,7 @@ let build_inversion_problem loc env sigma tms t = let patl,acc = List.fold_map' reveal_pattern realargs acc in let pat,acc = make_patvar t acc in let indf' = lift_inductive_family n indf in - let sign = make_arity_signature env true indf' in - let sign = List.map (fun d -> map_rel_decl EConstr.of_constr d) sign in + let sign = make_arity_signature env sigma true indf' in let patl = pat :: List.rev patl in let patl,sign = recover_and_adjust_alias_names patl sign in let p = List.length patl in diff --git a/pretyping/detyping.ml b/pretyping/detyping.ml index c8d945c0b..1adda14ab 100644 --- a/pretyping/detyping.ml +++ b/pretyping/detyping.ml @@ -6,14 +6,17 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +module CVars = Vars + open Pp open CErrors open Util open Names open Term +open Environ +open EConstr open Vars open Inductiveops -open Environ open Glob_term open Glob_ops open Termops @@ -188,7 +191,7 @@ let _ = declare_bool_option (* Auxiliary function for MutCase printing *) (* [computable] tries to tell if the predicate typing the result is inferable*) -let computable p k = +let computable sigma p k = (* We first remove as many lambda as the arity, then we look if it remains a lambda for a dependent elimination. This function works for normal eta-expanded term. For non eta-expanded or @@ -205,31 +208,29 @@ let computable p k = sinon on perd la réciprocité de la synthèse (qui, lui, engendrera un prédicat non dépendant) *) - let sign,ccl = decompose_lam_assum p in + let sign,ccl = decompose_lam_assum sigma p in Int.equal (Context.Rel.length sign) (k + 1) && - noccur_between 1 (k+1) ccl + noccur_between sigma 1 (k+1) ccl -let pop t = Vars.lift (-1) t - -let lookup_name_as_displayed env t s = - let rec lookup avoid n c = match kind_of_term c with +let lookup_name_as_displayed env sigma t s = + let rec lookup avoid n c = match EConstr.kind sigma c with | Prod (name,_,c') -> - (match compute_displayed_name_in RenamingForGoal avoid name c' with + (match compute_displayed_name_in sigma RenamingForGoal avoid name c' with | (Name id,avoid') -> if Id.equal id s then Some n else lookup avoid' (n+1) c' | (Anonymous,avoid') -> lookup avoid' (n+1) (pop c')) | LetIn (name,_,_,c') -> - (match compute_displayed_name_in RenamingForGoal avoid name c' with + (match compute_displayed_name_in sigma RenamingForGoal avoid name c' with | (Name id,avoid') -> if Id.equal id s then Some n else lookup avoid' (n+1) c' | (Anonymous,avoid') -> lookup avoid' (n+1) (pop c')) | Cast (c,_,_) -> lookup avoid n c | _ -> None in lookup (ids_of_named_context (named_context env)) 1 t -let lookup_index_as_renamed env t n = - let rec lookup n d c = match kind_of_term c with +let lookup_index_as_renamed env sigma t n = + let rec lookup n d c = match EConstr.kind sigma c with | Prod (name,_,c') -> - (match compute_displayed_name_in RenamingForGoal [] name c' with + (match compute_displayed_name_in sigma RenamingForGoal [] name c' with (Name _,_) -> lookup n (d+1) c' | (Anonymous,_) -> if Int.equal n 0 then @@ -239,7 +240,7 @@ let lookup_index_as_renamed env t n = else lookup (n-1) (d+1) c') | LetIn (name,_,_,c') -> - (match compute_displayed_name_in RenamingForGoal [] name c' with + (match compute_displayed_name_in sigma RenamingForGoal [] name c' with | (Name _,_) -> lookup n (d+1) c' | (Anonymous,_) -> if Int.equal n 0 then @@ -256,9 +257,9 @@ let lookup_index_as_renamed env t n = (**********************************************************************) (* Fragile algorithm to reverse pattern-matching compilation *) -let update_name na ((_,(e,_)),c) = +let update_name sigma na ((_,(e,_)),c) = match na with - | Name _ when force_wildcard () && noccurn (List.index Name.equal na e) c -> + | Name _ when force_wildcard () && noccurn sigma (List.index Name.equal na e) c -> Anonymous | _ -> na @@ -269,7 +270,7 @@ let rec decomp_branch tags nal b (avoid,env as e) sigma c = | [] -> (List.rev nal,(e,c)) | b::tags -> let na,c,f,body,t = - match kind_of_term (EConstr.Unsafe.to_constr (strip_outer_cast sigma (EConstr.of_constr c))), b with + match EConstr.kind sigma (strip_outer_cast sigma c), b with | Lambda (na,t,c),false -> na,c,compute_displayed_let_name_in,None,Some t | LetIn (na,b,t,c),true -> na,c,compute_displayed_name_in,Some b,Some t @@ -279,12 +280,12 @@ let rec decomp_branch tags nal b (avoid,env as e) sigma c = | _, true -> Anonymous,lift 1 c,compute_displayed_name_in,None,None in - let na',avoid' = f flag avoid na c in + let na',avoid' = f sigma flag avoid na c in decomp_branch tags (na'::nal) b (avoid', add_name_opt na' body t env) sigma c let rec build_tree na isgoal e sigma ci cl = - let mkpat n rhs pl = PatCstr(dl,(ci.ci_ind,n+1),pl,update_name na rhs) in + let mkpat n rhs pl = PatCstr(dl,(ci.ci_ind,n+1),pl,update_name sigma na rhs) in let cnl = ci.ci_pp_info.cstr_tags in let cna = ci.ci_cstr_nargs in List.flatten @@ -294,12 +295,12 @@ let rec build_tree na isgoal e sigma ci cl = and align_tree nal isgoal (e,c as rhs) sigma = match nal with | [] -> [[],rhs] | na::nal -> - match kind_of_term c with + match EConstr.kind sigma c with | Case (ci,p,c,cl) when - eq_constr sigma (EConstr.of_constr c) (EConstr.mkRel (List.index Name.equal na (fst (snd e)))) + eq_constr sigma c (mkRel (List.index Name.equal na (fst (snd e)))) && not (Int.equal (Array.length cl) 0) && (* don't contract if p dependent *) - computable p (List.length ci.ci_pp_info.ind_tags) (* FIXME: can do better *) -> + computable sigma p (List.length ci.ci_pp_info.ind_tags) (* FIXME: can do better *) -> let clauses = build_tree na isgoal e sigma ci cl in List.flatten (List.map (fun (pat,rhs) -> @@ -307,7 +308,7 @@ and align_tree nal isgoal (e,c as rhs) sigma = match nal with List.map (fun (hd,rest) -> pat::hd,rest) lines) clauses) | _ -> - let pat = PatVar(dl,update_name na rhs) in + let pat = PatVar(dl,update_name sigma na rhs) in let mat = align_tree nal isgoal rhs sigma in List.map (fun (hd,rest) -> pat::hd,rest) mat @@ -320,11 +321,11 @@ and contract_branch isgoal e sigma (cdn,can,mkpat,b) = (* Transform internal representation of pattern-matching into list of *) (* clauses *) -let is_nondep_branch c l = +let is_nondep_branch sigma c l = try (* FIXME: do better using tags from l *) - let sign,ccl = decompose_lam_n_decls (List.length l) c in - noccur_between 1 (Context.Rel.length sign) ccl + let sign,ccl = decompose_lam_n_decls sigma (List.length l) c in + noccur_between sigma 1 (Context.Rel.length sign) ccl with e when CErrors.noncritical e -> (* Not eta-expanded or not reduced *) false @@ -441,7 +442,7 @@ let detype_instance sigma l = else Some (List.map (detype_level sigma) (Array.to_list (Univ.Instance.to_array l))) let rec detype flags avoid env sigma t = - match kind_of_term (EConstr.Unsafe.to_constr (collapse_appl sigma (EConstr.of_constr t))) with + match EConstr.kind sigma (collapse_appl sigma t) with | Rel n -> (try match lookup_name_of_rel n (fst env) with | Name id -> GVar (dl, id) @@ -503,11 +504,11 @@ let rec detype flags avoid env sigma t = try let pb = Environ.lookup_projection p (snd env) in let body = pb.Declarations.proj_body in - let ty = Retyping.get_type_of (snd env) sigma (EConstr.of_constr c) in + let ty = Retyping.get_type_of (snd env) sigma c in let ((ind,u), args) = Inductiveops.find_mrectype (snd env) sigma ty in - let args = List.map EConstr.Unsafe.to_constr args in let body' = strip_lam_assum body in - let body' = subst_instance_constr u body' in + let body' = CVars.subst_instance_constr u body' in + let body' = EConstr.of_constr body' in substl (c :: List.rev args) body' with Retyping.RetypeError _ | Not_found -> anomaly (str"Cannot detype an unfolded primitive projection.") @@ -515,8 +516,7 @@ let rec detype flags avoid env sigma t = else if print_primproj_params () then try - let c = Retyping.expand_projection (snd env) sigma p (EConstr.of_constr c) [] in - let c = EConstr.Unsafe.to_constr c in + let c = Retyping.expand_projection (snd env) sigma p c [] in detype flags avoid env sigma c with Retyping.RetypeError _ -> noparams () else noparams () @@ -527,8 +527,8 @@ let rec detype flags avoid env sigma t = | LocalDef _ -> true | LocalAssum (id,_) -> try let n = List.index Name.equal (Name id) (fst env) in - isRelN n c - with Not_found -> isVarId id c + isRelN sigma n c + with Not_found -> isVarId sigma id c in let id,l = try @@ -537,8 +537,8 @@ let rec detype flags avoid env sigma t = | Some id -> id in let l = Evd.evar_instance_array bound_to_itself_or_letin (Evd.find sigma evk) cl in - let fvs,rels = List.fold_left (fun (fvs,rels) (_,c) -> match kind_of_term c with Rel n -> (fvs,Int.Set.add n rels) | Var id -> (Id.Set.add id fvs,rels) | _ -> (fvs,rels)) (Id.Set.empty,Int.Set.empty) l in - let l = Evd.evar_instance_array (fun d c -> not !print_evar_arguments && (bound_to_itself_or_letin d c && not (isRel c && Int.Set.mem (destRel c) rels || isVar c && (Id.Set.mem (destVar c) fvs)))) (Evd.find sigma evk) cl in + let fvs,rels = List.fold_left (fun (fvs,rels) (_,c) -> match EConstr.kind sigma c with Rel n -> (fvs,Int.Set.add n rels) | Var id -> (Id.Set.add id fvs,rels) | _ -> (fvs,rels)) (Id.Set.empty,Int.Set.empty) l in + let l = Evd.evar_instance_array (fun d c -> not !print_evar_arguments && (bound_to_itself_or_letin d c && not (isRel sigma c && Int.Set.mem (destRel sigma c) rels || isVar sigma c && (Id.Set.mem (destVar sigma c) fvs)))) (Evd.find sigma evk) cl in id,l with Not_found -> Id.of_string ("X" ^ string_of_int (Evar.repr evk)), @@ -551,10 +551,10 @@ let rec detype flags avoid env sigma t = | Construct (cstr_sp,u) -> GRef (dl, ConstructRef cstr_sp, detype_instance sigma u) | Case (ci,p,c,bl) -> - let comp = computable p (List.length (ci.ci_pp_info.ind_tags)) in + let comp = computable sigma p (List.length (ci.ci_pp_info.ind_tags)) in detype_case comp (detype flags avoid env sigma) (detype_eqns flags avoid env sigma ci comp) - is_nondep_branch avoid + (is_nondep_branch sigma) avoid (ci.ci_ind,ci.ci_pp_info.style, ci.ci_pp_info.cstr_tags,ci.ci_pp_info.ind_tags) (Some p) c bl @@ -594,7 +594,7 @@ and detype_cofix flags avoid env sigma n (names,tys,bodies) = Array.map (fun (_,bd,_) -> bd) v) and share_names flags n l avoid env sigma c t = - match kind_of_term c, kind_of_term t with + match EConstr.kind sigma c, EConstr.kind sigma t with (* factorize even when not necessary to have better presentation *) | Lambda (na,t,c), Prod (na',t',c') -> let na = match (na,na') with @@ -641,15 +641,15 @@ and detype_eqns flags avoid env sigma ci computable constructs consnargsl bl = and detype_eqn (lax,isgoal as flags) avoid env sigma constr construct_nargs branch = let make_pat x avoid env b body ty ids = - if force_wildcard () && noccurn 1 b then + if force_wildcard () && noccurn sigma 1 b then PatVar (dl,Anonymous),avoid,(add_name Anonymous body ty env),ids else let flag = if isgoal then RenamingForGoal else RenamingForCasesPattern (fst env,b) in - let na,avoid' = compute_displayed_name_in flag avoid x b in + let na,avoid' = compute_displayed_name_in sigma flag avoid x b in PatVar (dl,na),avoid',(add_name na body ty env),add_vname ids na in let rec buildrec ids patlist avoid env l b = - match kind_of_term b, l with + match EConstr.kind sigma b, l with | _, [] -> (dl, Id.Set.elements ids, [PatCstr(dl, constr, List.rev patlist,Anonymous)], @@ -684,8 +684,8 @@ and detype_eqn (lax,isgoal as flags) avoid env sigma constr construct_nargs bran and detype_binder (lax,isgoal as flags) bk avoid env sigma na body ty c = let flag = if isgoal then RenamingForGoal else RenamingElsewhereFor (fst env,c) in let na',avoid' = match bk with - | BLetIn -> compute_displayed_let_name_in flag avoid na c - | _ -> compute_displayed_name_in flag avoid na c in + | BLetIn -> compute_displayed_let_name_in sigma flag avoid na c + | _ -> compute_displayed_name_in sigma flag avoid na c in let r = detype flags avoid' (add_name na' body ty env) sigma c in match bk with | BProd -> GProd (dl, na',Explicit,detype (lax,false) avoid env sigma ty, r) @@ -693,13 +693,13 @@ and detype_binder (lax,isgoal as flags) bk avoid env sigma na body ty c = | BLetIn -> let c = detype (lax,false) avoid env sigma (Option.get body) in (* Heuristic: we display the type if in Prop *) - let s = try Retyping.get_sort_family_of (snd env) sigma (EConstr.of_constr ty) with _ when !Flags.in_debugger || !Flags.in_toplevel -> InType (* Can fail because of sigma missing in debugger *) in + let s = try Retyping.get_sort_family_of (snd env) sigma ty with _ when !Flags.in_debugger || !Flags.in_toplevel -> InType (* Can fail because of sigma missing in debugger *) in let c = if s != InProp then c else GCast (dl, c, CastConv (detype (lax,false) avoid env sigma ty)) in GLetIn (dl, na', c, r) let detype_rel_context ?(lax=false) where avoid env sigma sign = - let where = Option.map (fun c -> it_mkLambda_or_LetIn c sign) where in + let where = Option.map (fun c -> EConstr.it_mkLambda_or_LetIn c sign) where in let rec aux avoid env = function | [] -> [] | decl::rest -> @@ -711,10 +711,10 @@ let detype_rel_context ?(lax=false) where avoid env sigma sign = | None -> na,avoid | Some c -> if is_local_def decl then - compute_displayed_let_name_in + compute_displayed_let_name_in sigma (RenamingElsewhereFor (fst env,c)) avoid na c else - compute_displayed_name_in + compute_displayed_name_in sigma (RenamingElsewhereFor (fst env,c)) avoid na c in let b = match decl with | LocalAssum _ -> None @@ -731,6 +731,7 @@ let detype ?(lax=false) isgoal avoid env sigma t = detype (lax,isgoal) avoid (names_of_rel_context env, env) sigma t let detype_closed_glob ?lax isgoal avoid env sigma t = + let open Context.Rel.Declaration in let convert_id cl id = try Id.Map.find id cl.idents with Not_found -> id @@ -748,12 +749,11 @@ let detype_closed_glob ?lax isgoal avoid env sigma t = with Not_found -> try (* assumes [detype] does not raise [Not_found] exceptions *) let (b,c) = Id.Map.find id cl.typed in - let c = EConstr.Unsafe.to_constr c in (* spiwack: I'm not sure it is the right thing to do, but I'm computing the detyping environment like [Printer.pr_constr_under_binders_env] does. *) - let assums = List.map (fun id -> (Name id,(* dummy *) mkProp)) b in - let env = Termops.push_rels_assum assums env in + let assums = List.map (fun id -> LocalAssum (Name id,(* dummy *) mkProp)) b in + let env = push_rel_context assums env in detype ?lax isgoal avoid env sigma c (* if [id] is bound to a [closed_glob_constr]. *) with Not_found -> try @@ -808,7 +808,7 @@ let rec subst_glob_constr subst raw = | GRef (loc,ref,u) -> let ref',t = subst_global subst ref in if ref' == ref then raw else - detype false [] (Global.env()) Evd.empty t + detype false [] (Global.env()) Evd.empty (EConstr.of_constr t) | GVar _ -> raw | GEvar _ -> raw diff --git a/pretyping/detyping.mli b/pretyping/detyping.mli index c51cb0f44..4c6f9129f 100644 --- a/pretyping/detyping.mli +++ b/pretyping/detyping.mli @@ -9,6 +9,7 @@ open Names open Term open Environ +open EConstr open Glob_term open Termops open Mod_subst @@ -45,13 +46,13 @@ val detype_case : val detype_sort : evar_map -> sorts -> glob_sort val detype_rel_context : ?lax:bool -> constr option -> Id.t list -> (names_context * env) -> - evar_map -> Context.Rel.t -> glob_decl list + evar_map -> rel_context -> glob_decl list val detype_closed_glob : ?lax:bool -> bool -> Id.t list -> env -> evar_map -> closed_glob_constr -> glob_constr (** look for the index of a named var or a nondep var as it is renamed *) -val lookup_name_as_displayed : env -> constr -> Id.t -> int option -val lookup_index_as_renamed : env -> constr -> int -> int option +val lookup_name_as_displayed : env -> evar_map -> constr -> Id.t -> int option +val lookup_index_as_renamed : env -> evar_map -> constr -> int -> int option val set_detype_anonymous : (Loc.t -> int -> glob_constr) -> unit val force_wildcard : unit -> bool diff --git a/pretyping/evardefine.ml b/pretyping/evardefine.ml index faf34baf7..20d86f81b 100644 --- a/pretyping/evardefine.ml +++ b/pretyping/evardefine.ml @@ -140,7 +140,7 @@ let define_pure_evar_as_lambda env evd evk = | _ -> error_not_product env evd typ in let avoid = ids_of_named_context (evar_context evi) in let id = - next_name_away_with_default_using_types "x" na avoid (EConstr.Unsafe.to_constr (Reductionops.whd_evar evd dom)) in + next_name_away_with_default_using_types "x" na avoid (Reductionops.whd_evar evd dom) in let newenv = push_named (LocalAssum (id, dom)) evenv in let filter = Filter.extend 1 (evar_filter evi) in let src = evar_source evk evd1 in diff --git a/pretyping/indrec.ml b/pretyping/indrec.ml index 431d1ff16..c4a74d990 100644 --- a/pretyping/indrec.ml +++ b/pretyping/indrec.ml @@ -40,6 +40,18 @@ type recursion_scheme_error = exception RecursionSchemeError of recursion_scheme_error +let named_hd env t na = named_hd env Evd.empty (EConstr.of_constr t) na +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 mkLambda_or_LetIn_name env d b = mkLambda_or_LetIn (name_assumption env d) b +let mkProd_or_LetIn_name env d b = mkProd_or_LetIn (name_assumption env d) b +let mkLambda_name env (n,a,b) = mkLambda_or_LetIn_name env (LocalAssum (n,a)) b +let mkProd_name env (n,a,b) = mkProd_or_LetIn_name env (LocalAssum (n,a)) b +let it_mkProd_or_LetIn_name env b l = List.fold_left (fun c d -> mkProd_or_LetIn_name env d c) b l +let it_mkLambda_or_LetIn_name env b l = List.fold_left (fun c d -> mkLambda_or_LetIn_name env d c) b l + let make_prod_dep dep env = if dep then mkProd_name env else mkProd let mkLambda_string s t c = mkLambda (Name (Id.of_string s), t, c) @@ -118,12 +130,13 @@ let mis_make_case_com dep env sigma (ind, u as pind) (mib,mip as specif) kind = it_mkLambda_or_LetIn_name env' obj deparsign else let cs = lift_constructor (k+1) constrs.(k) in - let t = build_branch_type env dep (mkRel (k+1)) cs in + let t = build_branch_type env (Sigma.to_evar_map sigma) dep (mkRel (k+1)) cs in mkLambda_string "f" t (add_branch (push_rel (LocalAssum (Anonymous, t)) env) (k+1)) in let Sigma (s, sigma, p) = Sigma.fresh_sort_in_family ~rigid:Evd.univ_flexible_alg env sigma kind in - let typP = make_arity env' dep indf s in + let typP = make_arity env' (Sigma.to_evar_map sigma) dep indf s in + let typP = EConstr.Unsafe.to_constr typP in let c = it_mkLambda_or_LetIn_name env (mkLambda_string "P" typP @@ -443,7 +456,8 @@ let mis_make_indrec env sigma listdepkind mib u = Evarutil.evd_comb1 (Evd.fresh_sort_in_family ~rigid:Evd.univ_flexible_alg env) evdref kinds in - let typP = make_arity env dep indf s in + let typP = make_arity env !evdref dep indf s in + let typP = EConstr.Unsafe.to_constr typP in mkLambda_string "P" typP (put_arity (push_rel (LocalAssum (Anonymous,typP)) env) (i+1) rest) | [] -> diff --git a/pretyping/inductiveops.ml b/pretyping/inductiveops.ml index 3191a58ff..9e823ab4c 100644 --- a/pretyping/inductiveops.ml +++ b/pretyping/inductiveops.ml @@ -422,26 +422,29 @@ let build_dependent_inductive env ((ind, params) as indf) = (* builds the arity of an elimination predicate in sort [s] *) -let make_arity_signature env dep indf = +let make_arity_signature env sigma dep indf = let (arsign,_) = get_arity env indf in + let arsign = List.map (fun d -> Termops.map_rel_decl EConstr.of_constr d) arsign in if dep then (* We need names everywhere *) - Namegen.name_context env - ((LocalAssum (Anonymous,build_dependent_inductive env indf))::arsign) + Namegen.name_context env sigma + ((LocalAssum (Anonymous,EConstr.of_constr (build_dependent_inductive env indf)))::arsign) (* Costly: would be better to name once for all at definition time *) else (* No need to enforce names *) arsign -let make_arity env dep indf s = mkArity (make_arity_signature env dep indf, s) +let make_arity env sigma dep indf s = + let open EConstr in + it_mkProd_or_LetIn (mkSort s) (make_arity_signature env sigma dep indf) (* [p] is the predicate and [cs] a constructor summary *) -let build_branch_type env dep p cs = +let build_branch_type env sigma dep p cs = let base = appvect (lift cs.cs_nargs p, cs.cs_concl_realargs) in if dep then - Namegen.it_mkProd_or_LetIn_name env - (applist (base,[build_dependent_constructor cs])) - cs.cs_args + EConstr.Unsafe.to_constr (Namegen.it_mkProd_or_LetIn_name env sigma + (EConstr.of_constr (applist (base,[build_dependent_constructor cs]))) + (List.map (fun d -> Termops.map_rel_decl EConstr.of_constr d) cs.cs_args)) else Term.it_mkProd_or_LetIn base cs.cs_args @@ -542,11 +545,12 @@ let is_elim_predicate_explicitly_dependent env sigma pred indf = let arsign,_ = get_arity env indf in is_predicate_explicitly_dep env sigma pred arsign -let set_names env n brty = - let (ctxt,cl) = decompose_prod_n_assum n brty in - Namegen.it_mkProd_or_LetIn_name env cl ctxt +let set_names env sigma n brty = + let open EConstr in + let (ctxt,cl) = decompose_prod_n_assum sigma n brty in + EConstr.Unsafe.to_constr (Namegen.it_mkProd_or_LetIn_name env sigma cl ctxt) -let set_pattern_names env ind brv = +let set_pattern_names env sigma ind brv = let (mib,mip) = Inductive.lookup_mind_specif env ind in let arities = Array.map @@ -554,7 +558,7 @@ let set_pattern_names env ind brv = Context.Rel.length ((prod_assum c)) - mib.mind_nparams) mip.mind_nf_lc in - Array.map2 (set_names env) arities brv + Array.map2 (set_names env sigma) arities brv let type_case_branches_with_names env sigma indspec p c = let (ind,args) = indspec in @@ -567,7 +571,7 @@ let type_case_branches_with_names env sigma indspec p c = let conclty = lambda_appvect_assum (mip.mind_nrealdecls+1) p (Array.of_list (realargs@[c])) in (* Adjust names *) if is_elim_predicate_explicitly_dependent env sigma p (ind,params) then - (set_pattern_names env (fst ind) lbrty, conclty) + (set_pattern_names env sigma (fst ind) (Array.map EConstr.of_constr lbrty), conclty) else (lbrty, conclty) (* Type of Case predicates *) diff --git a/pretyping/inductiveops.mli b/pretyping/inductiveops.mli index 4bb484759..ab470a540 100644 --- a/pretyping/inductiveops.mli +++ b/pretyping/inductiveops.mli @@ -156,9 +156,9 @@ val get_arity : env -> inductive_family -> Context.Rel.t * sorts_family val build_dependent_constructor : constructor_summary -> constr val build_dependent_inductive : env -> inductive_family -> constr -val make_arity_signature : env -> bool -> inductive_family -> Context.Rel.t -val make_arity : env -> bool -> inductive_family -> sorts -> types -val build_branch_type : env -> bool -> constr -> constructor_summary -> types +val make_arity_signature : env -> evar_map -> bool -> inductive_family -> EConstr.rel_context +val make_arity : env -> evar_map -> bool -> inductive_family -> sorts -> EConstr.types +val build_branch_type : env -> evar_map -> bool -> constr -> constructor_summary -> types (** Raise [Not_found] if not given a valid inductive type *) val extract_mrectype : evar_map -> EConstr.t -> pinductive * EConstr.constr list diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml index 2470decdd..563769df5 100644 --- a/pretyping/pretyping.ml +++ b/pretyping/pretyping.ml @@ -79,26 +79,26 @@ type t = { (** Delay the computation of the evar extended environment *) } -let get_extra env = +let get_extra env sigma = let open Context.Named.Declaration in let ids = List.map get_id (named_context env) in let avoid = List.fold_right Id.Set.add ids Id.Set.empty in - 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) -let make_env env = { env = env; extra = lazy (get_extra env) } +let make_env env sigma = { env = env; extra = lazy (get_extra env sigma) } let rel_context env = rel_context env.env -let push_rel d env = { +let push_rel sigma d env = { env = push_rel d env.env; - extra = lazy (push_rel_decl_to_named_context d (Lazy.force env.extra)); + extra = lazy (push_rel_decl_to_named_context sigma d (Lazy.force env.extra)); } -let pop_rel_context n env = make_env (pop_rel_context n env.env) +let pop_rel_context n env sigma = make_env (pop_rel_context n env.env) sigma -let push_rel_context ctx env = { +let push_rel_context sigma ctx env = { env = push_rel_context ctx env.env; - extra = lazy (List.fold_right push_rel_decl_to_named_context ctx (Lazy.force env.extra)); + extra = lazy (List.fold_right (fun d acc -> push_rel_decl_to_named_context sigma d acc) ctx (Lazy.force env.extra)); } let lookup_named id env = lookup_named id env.env @@ -117,9 +117,9 @@ let e_new_evar env evdref ?src ?naming typ = evdref := Sigma.to_evar_map sigma; e -let push_rec_types (lna,typarray,_) env = +let push_rec_types sigma (lna,typarray,_) env = let ctxt = Array.map2_i (fun i na t -> Context.Rel.Declaration.LocalAssum (na, lift i t)) lna typarray in - Array.fold_left (fun e assum -> push_rel assum e) env ctxt + Array.fold_left (fun e assum -> push_rel sigma assum e) env ctxt end @@ -391,7 +391,7 @@ let ltac_interp_name { ltac_idents ; ltac_genargs } = function str"It cannot be used in a binder.") else n -let ltac_interp_name_env k0 lvar env = +let ltac_interp_name_env k0 lvar env sigma = (* envhd is the initial part of the env when pretype was called first *) (* (in practice is is probably 0, but we have to grant the specification of pretype which accepts to start with a non empty @@ -402,7 +402,7 @@ let ltac_interp_name_env k0 lvar env = let open Context.Rel.Declaration in let ctxt' = List.smartmap (map_name (ltac_interp_name lvar)) ctxt in if List.equal (fun d1 d2 -> Name.equal (get_name d1) (get_name d2)) ctxt ctxt' then env - else push_rel_context ctxt' (pop_rel_context n env) + else push_rel_context sigma ctxt' (pop_rel_context n env sigma) let invert_ltac_bound_name lvar env id0 id = let id' = Id.Map.find id lvar.ltac_idents in @@ -426,7 +426,7 @@ let pretype_id pretype k0 loc env evdref lvar id = let (n,_,typ) = lookup_rel_id id (rel_context env) in { uj_val = mkRel n; uj_type = lift n typ } with Not_found -> - let env = ltac_interp_name_env k0 lvar env in + let env = ltac_interp_name_env k0 lvar env !evdref in (* Check if [id] is an ltac variable *) try let (ids,c) = Id.Map.find id lvar.ltac_constrs in @@ -569,7 +569,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre inh_conv_coerce_to_tycon loc env evdref j tycon | GPatVar (loc,(someta,n)) -> - let env = ltac_interp_name_env k0 lvar env in + let env = ltac_interp_name_env k0 lvar env !evdref in let ty = match tycon with | Some ty -> ty @@ -578,7 +578,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre { uj_val = e_new_evar env evdref ~src:(loc,k) ty; uj_type = ty } | GHole (loc, k, naming, None) -> - let env = ltac_interp_name_env k0 lvar env in + let env = ltac_interp_name_env k0 lvar env !evdref in let ty = match tycon with | Some ty -> ty @@ -587,7 +587,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre { uj_val = e_new_evar env evdref ~src:(loc,k) ~naming ty; uj_type = ty } | GHole (loc, k, _naming, Some arg) -> - let env = ltac_interp_name_env k0 lvar env in + let env = ltac_interp_name_env k0 lvar env !evdref in let ty = match tycon with | Some ty -> ty @@ -605,18 +605,18 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre let ty' = pretype_type empty_valcon env evdref lvar ty in let dcl = LocalAssum (na, ty'.utj_val) in let dcl' = LocalAssum (ltac_interp_name lvar na,ty'.utj_val) in - type_bl (push_rel dcl env) (Context.Rel.add dcl' ctxt) bl + type_bl (push_rel !evdref dcl env) (Context.Rel.add dcl' ctxt) bl | (na,bk,Some bd,ty)::bl -> let ty' = pretype_type empty_valcon env evdref lvar ty in let bd' = pretype (mk_tycon ty'.utj_val) env evdref lvar bd in let dcl = LocalDef (na, bd'.uj_val, ty'.utj_val) in let dcl' = LocalDef (ltac_interp_name lvar na, bd'.uj_val, ty'.utj_val) in - type_bl (push_rel dcl env) (Context.Rel.add dcl' ctxt) bl in + type_bl (push_rel !evdref dcl env) (Context.Rel.add dcl' ctxt) bl in let ctxtv = Array.map (type_bl env Context.Rel.empty) bl in let larj = Array.map2 (fun e ar -> - pretype_type empty_valcon (push_rel_context e env) evdref lvar ar) + pretype_type empty_valcon (push_rel_context !evdref e env) evdref lvar ar) ctxtv lar in let lara = Array.map (fun a -> a.utj_val) larj in let ftys = Array.map2 (fun e a -> it_mkProd_or_LetIn a e) ctxtv lara in @@ -632,7 +632,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre | None -> true in (* Note: bodies are not used by push_rec_types, so [||] is safe *) - let newenv = push_rec_types (names,ftys,[||]) env in + let newenv = push_rec_types !evdref (names,ftys,[||]) env in let vdefj = Array.map2_i (fun i ctxt def -> @@ -641,7 +641,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre let (ctxt,ty) = decompose_prod_n_assum !evdref (Context.Rel.length ctxt) (lift nbfix ftys.(i)) in - let nenv = push_rel_context ctxt newenv in + let nenv = push_rel_context !evdref ctxt newenv in let j = pretype (mk_tycon ty) nenv evdref lvar def in { uj_val = it_mkLambda_or_LetIn j.uj_val ctxt; uj_type = it_mkProd_or_LetIn j.uj_type ctxt }) @@ -783,7 +783,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre the substitution must also be applied on variables before they are looked up in the rel context. *) let var = LocalAssum (name, j.utj_val) in - let j' = pretype rng (push_rel var env) evdref lvar c2 in + let j' = pretype rng (push_rel !evdref var env) evdref lvar c2 in let name = ltac_interp_name lvar name in let resj = judge_of_abstraction env.ExtraEnv.env (orelse_name name name') j j' in inh_conv_coerce_to_tycon loc env evdref resj tycon @@ -799,7 +799,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre { j with utj_val = lift 1 j.utj_val } | Name _ -> let var = LocalAssum (name, j.utj_val) in - let env' = push_rel var env in + let env' = push_rel !evdref var env in pretype_type empty_valcon env' evdref lvar c2 in let name = ltac_interp_name lvar name in @@ -828,7 +828,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre looked up in the rel context. *) let var = LocalDef (name, j.uj_val, t) in let tycon = lift_tycon 1 tycon in - let j' = pretype tycon (push_rel var env) evdref lvar c2 in + let j' = pretype tycon (push_rel !evdref var env) evdref lvar c2 in let name = ltac_interp_name lvar name in { uj_val = mkLetIn (name, j.uj_val, t, j'.uj_val) ; uj_type = subst1 j.uj_val j'.uj_type } @@ -877,7 +877,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre mkCase (ci, p, cj.uj_val,[|f|]) else it_mkLambda_or_LetIn f fsign in - let env_f = push_rel_context fsign env in + let env_f = push_rel_context !evdref fsign env in (* Make dependencies from arity signature impossible *) let arsgn = let arsgn,_ = get_arity env.ExtraEnv.env indf in @@ -890,11 +890,10 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre let nar = List.length arsgn in (match po with | Some p -> - let env_p = push_rel_context psign env in + let env_p = push_rel_context !evdref psign env in let pj = pretype_type empty_valcon env_p evdref lvar p in let ccl = nf_evar !evdref pj.utj_val in - let psign = make_arity_signature env.ExtraEnv.env true indf in (* with names *) - let psign = List.map (fun d -> map_rel_decl EConstr.of_constr d) psign in + let psign = make_arity_signature env.ExtraEnv.env !evdref true indf in (* with names *) let p = it_mkLambda_or_LetIn ccl psign in let inst = (Array.map_to_list EConstr.of_constr cs.cs_concl_realargs) @@ -951,7 +950,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre let psign = List.map (fun d -> map_rel_decl EConstr.of_constr d) psign in let pred,p = match po with | Some p -> - let env_p = push_rel_context psign env in + let env_p = push_rel_context !evdref psign env in let pj = pretype_type empty_valcon env_p evdref lvar p in let ccl = nf_evar !evdref pj.utj_val in let pred = it_mkLambda_or_LetIn ccl psign in @@ -961,7 +960,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre let p = match tycon with | Some ty -> ty | None -> - let env = ltac_interp_name_env k0 lvar env in + let env = ltac_interp_name_env k0 lvar env !evdref in new_type_evar env evdref loc in it_mkLambda_or_LetIn (lift (nar+1) p) psign, p in @@ -980,7 +979,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre | Anonymous -> Name Namegen.default_non_dependent_ident)) cs_args in - let env_c = push_rel_context csgn env in + let env_c = push_rel_context !evdref csgn env in let bj = pretype (mk_tycon pi) env_c evdref lvar b in it_mkLambda_or_LetIn bj.uj_val cs_args in let b1 = f cstrs.(0) b1 in @@ -997,7 +996,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre | GCases (loc,sty,po,tml,eqns) -> Cases.compile_cases loc sty - ((fun vtyc env evdref -> pretype vtyc (make_env env) evdref lvar),evdref) + ((fun vtyc env evdref -> pretype vtyc (make_env env !evdref) evdref lvar),evdref) tycon env.ExtraEnv.env (* loc *) (po,tml,eqns) | GCast (loc,c,k) -> @@ -1090,7 +1089,7 @@ and pretype_type k0 resolve_tc valcon (env : ExtraEnv.t) evdref lvar = function { utj_val = v; utj_type = s } | None -> - let env = ltac_interp_name_env k0 lvar env in + let env = ltac_interp_name_env k0 lvar env !evdref in let s = evd_comb0 (new_sort_variable univ_flexible_alg) evdref in { utj_val = e_new_evar env evdref ~src:(loc, knd) ~naming (mkSort s); utj_type = s}) @@ -1107,7 +1106,7 @@ and pretype_type k0 resolve_tc valcon (env : ExtraEnv.t) evdref lvar = function ~loc:(loc_of_glob_constr c) env.ExtraEnv.env !evdref tj.utj_val v let ise_pretype_gen flags env sigma lvar kind c = - let env = make_env env in + let env = make_env env sigma in let evdref = ref sigma in let k0 = Context.Rel.length (rel_context env) in let c' = match kind with @@ -1150,7 +1149,7 @@ let on_judgment sigma f j = {uj_val = c; uj_type = t} let understand_judgment env sigma c = - let env = make_env env in + let env = make_env env sigma in let evdref = ref sigma in let k0 = Context.Rel.length (rel_context env) in let j = pretype k0 true empty_tycon env evdref empty_lvar c in @@ -1160,7 +1159,7 @@ let understand_judgment env sigma c = in j, Evd.evar_universe_context !evdref let understand_judgment_tcc env evdref c = - let env = make_env env in + let env = make_env env !evdref in let k0 = Context.Rel.length (rel_context env) in let j = pretype k0 true empty_tycon env evdref empty_lvar c in on_judgment !evdref (fun c -> @@ -1217,7 +1216,7 @@ let type_uconstr ?(flags = constr_flags) end } let pretype k0 resolve_tc typcon env evdref lvar t = - pretype k0 resolve_tc typcon (make_env env) evdref lvar t + pretype k0 resolve_tc typcon (make_env env !evdref) evdref lvar t let pretype_type k0 resolve_tc valcon env evdref lvar t = - pretype_type k0 resolve_tc valcon (make_env env) evdref lvar t + pretype_type k0 resolve_tc valcon (make_env env !evdref) evdref lvar t diff --git a/pretyping/tacred.ml b/pretyping/tacred.ml index 9f3f3c7e5..ef9f39d77 100644 --- a/pretyping/tacred.ml +++ b/pretyping/tacred.ml @@ -1141,7 +1141,7 @@ let compute = cbv_betadeltaiota let abstract_scheme env sigma (locc,a) (c, sigma) = let ta = Retyping.get_type_of env sigma a in - let na = named_hd env (EConstr.to_constr sigma ta) Anonymous in + let na = named_hd env sigma ta Anonymous in if occur_meta sigma ta then error "Cannot find a type for the generalisation."; if occur_meta sigma a then mkLambda (na,ta,c), sigma diff --git a/pretyping/typing.ml b/pretyping/typing.ml index bdd3663d1..dec22ecd0 100644 --- a/pretyping/typing.ml +++ b/pretyping/typing.ml @@ -92,8 +92,7 @@ let max_sort l = if Sorts.List.mem InSet l then InSet else InProp let e_is_correct_arity env evdref c pj ind specif params = - let arsign = make_arity_signature env true (make_ind_family (ind,params)) in - let arsign = List.map (fun d -> Termops.map_rel_decl EConstr.of_constr d) arsign in + let arsign = make_arity_signature env !evdref true (make_ind_family (ind,params)) in let allowed_sorts = elim_sorts specif in let error () = Pretype_errors.error_elim_arity env !evdref ind allowed_sorts c pj None in let rec srec env pt ar = diff --git a/pretyping/unification.ml b/pretyping/unification.ml index 589201fe2..baa12db08 100644 --- a/pretyping/unification.ml +++ b/pretyping/unification.ml @@ -85,7 +85,7 @@ let occur_meta_evd sigma mv c = let abstract_scheme env evd c l lname_typ = let mkLambda_name env (n,a,b) = - mkLambda (named_hd env (EConstr.Unsafe.to_constr a) n, a, b) + mkLambda (named_hd env evd a n, a, b) in List.fold_left2 (fun (t,evd) (locc,a) decl -> @@ -1617,8 +1617,7 @@ let make_eq_test env evd c = let make_abstraction_core name (test,out) env sigma c ty occs check_occs concl = let id = let t = match ty with Some t -> t | None -> get_type_of env sigma c in - let t = EConstr.Unsafe.to_constr t in - let x = id_of_name_using_hdchar (Global.env()) t name in + let x = id_of_name_using_hdchar (Global.env()) sigma t name in let ids = ids_of_named_context (named_context env) in if name == Anonymous then next_ident_away_in_goal x ids else if mem_named_context_val x (named_context_val env) then diff --git a/printing/printer.ml b/printing/printer.ml index c7d3a1ba3..447337b9a 100644 --- a/printing/printer.ml +++ b/printing/printer.ml @@ -515,7 +515,7 @@ let print_evar_constraints gl sigma = 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 + Namegen.make_all_name_different env sigma in str" " ++ hov 2 (pr_env env ++ pr_leconstr_env env sigma t1 ++ spc () ++ str (match pbty with diff --git a/proofs/clenv.ml b/proofs/clenv.ml index e580bccc3..5645850b2 100644 --- a/proofs/clenv.ml +++ b/proofs/clenv.ml @@ -555,7 +555,7 @@ let make_clenv_binding_gen hyps_only n env sigma (c,t) = function let clause = mk_clenv_from_env env sigma n (c,t) in clenv_constrain_dep_args hyps_only largs clause | ExplicitBindings lbind -> - let t = EConstr.of_constr (rename_bound_vars_as_displayed [] [] (EConstr.Unsafe.to_constr t)) in + let t = rename_bound_vars_as_displayed sigma [] [] t in let clause = mk_clenv_from_env env sigma n (c, t) in clenv_match_args lbind clause @@ -603,9 +603,7 @@ let make_evar_clause env sigma ?len t = | Some n -> assert (0 <= n); n in (** FIXME: do the renaming online *) - let t = EConstr.Unsafe.to_constr t in - let t = rename_bound_vars_as_displayed [] [] t in - let t = EConstr.of_constr t in + let t = rename_bound_vars_as_displayed sigma [] [] t in let rec clrec (sigma, holes) n t = if n = 0 then (sigma, holes, t) else match EConstr.kind sigma t with diff --git a/tactics/class_tactics.ml b/tactics/class_tactics.ml index 55fda1c7d..669d80814 100644 --- a/tactics/class_tactics.ml +++ b/tactics/class_tactics.ml @@ -1458,7 +1458,7 @@ let _ = Hook.set Typeclasses.solve_all_instances_hook solve_inst let resolve_one_typeclass env ?(sigma=Evd.empty) gl unique = - let nc, gl, subst, _, _ = Evarutil.push_rel_context_to_named_context env gl in + let nc, gl, subst, _, _ = Evarutil.push_rel_context_to_named_context env sigma gl in let (gl,t,sigma) = Goal.V82.mk_goal sigma nc gl Store.empty in let (ev, _) = destEvar sigma t in diff --git a/tactics/eqschemes.ml b/tactics/eqschemes.ml index 188e215a5..b08456f2f 100644 --- a/tactics/eqschemes.ml +++ b/tactics/eqschemes.ml @@ -76,10 +76,24 @@ let build_dependent_inductive ind (mib,mip) = Context.Rel.to_extended_list mkRel mip.mind_nrealdecls mib.mind_params_ctxt @ Context.Rel.to_extended_list mkRel 0 realargs) +let named_hd env t na = named_hd env Evd.empty (EConstr.of_constr t) na +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_context env hyps = + snd + (List.fold_left + (fun (env,hyps) d -> + let d' = name_assumption env d in (push_rel d' env, d' :: hyps)) + (env,[]) (List.rev hyps)) + let my_it_mkLambda_or_LetIn s c = it_mkLambda_or_LetIn c s let my_it_mkProd_or_LetIn s c = Term.it_mkProd_or_LetIn c s let my_it_mkLambda_or_LetIn_name s c = - it_mkLambda_or_LetIn_name (Global.env()) c s + let env = Global.env () in + let mkLambda_or_LetIn_name d b = mkLambda_or_LetIn (name_assumption env d) b in + List.fold_left (fun c d -> mkLambda_or_LetIn_name d c) c s let get_coq_eq ctx = try diff --git a/tactics/equality.ml b/tactics/equality.ml index 7f7a07b8f..d9b668517 100644 --- a/tactics/equality.ml +++ b/tactics/equality.ml @@ -857,14 +857,13 @@ let descend_then env sigma head dirn = (dirn_nlams, dirn_env, (fun dirnval (dfltval,resty) -> - let deparsign = make_arity_signature env true indf in - let deparsign = List.map (fun d -> map_rel_decl EConstr.of_constr d) deparsign in + let deparsign = make_arity_signature env sigma true indf in let p = it_mkLambda_or_LetIn (lift (mip.mind_nrealargs+1) resty) deparsign in let build_branch i = let result = if Int.equal i dirn then dirnval else dfltval in - let args = name_context env cstr.(i-1).cs_args in - let args = List.map (fun d -> map_rel_decl EConstr.of_constr d) args in + let cs_args = List.map (fun d -> map_rel_decl EConstr.of_constr d) cstr.(i-1).cs_args in + let args = name_context env sigma cs_args in it_mkLambda_or_LetIn result args in let brl = List.map build_branch @@ -905,8 +904,7 @@ let build_selector env sigma dirn c ind special default = let ind, _ = check_privacy env indp in let typ = Retyping.get_type_of env sigma default in let (mib,mip) = lookup_mind_specif env ind in - let deparsign = make_arity_signature env true indf in - let deparsign = List.map (fun d -> map_rel_decl EConstr.of_constr d) deparsign in + let deparsign = make_arity_signature env sigma true indf in let p = it_mkLambda_or_LetIn typ deparsign in let cstrs = get_constructors env indf in let build_branch i = @@ -1535,9 +1533,6 @@ let decomp_tuple_term env sigma c t = in [((ex,exty),inner_code)]::iterated_decomp in decomprec (mkRel 1) c t -let lambda_create env (a,b) = - mkLambda (named_hd env (EConstr.Unsafe.to_constr a) Anonymous, a, b) - let subst_tuple_term env sigma dep_pair1 dep_pair2 b = let sigma = Sigma.to_evar_map sigma in let typ = get_type_of env sigma dep_pair1 in @@ -1555,9 +1550,9 @@ let subst_tuple_term env sigma dep_pair1 dep_pair2 b = (* We build the expected goal *) let abst_B = List.fold_right - (fun (e,t) body -> lambda_create env (t,subst_term sigma e body)) e1_list b in + (fun (e,t) body -> lambda_create env sigma (t,subst_term sigma e body)) e1_list b in let pred_body = beta_applist sigma (abst_B,proj_list) in - let body = mkApp (lambda_create env (typ,pred_body),[|dep_pair1|]) in + let body = mkApp (lambda_create env sigma (typ,pred_body),[|dep_pair1|]) in let expected_goal = beta_applist sigma (abst_B,List.map fst e2_list) in (* Simulate now the normalisation treatment made by Logic.mk_refgoals *) let expected_goal = nf_betaiota sigma expected_goal in diff --git a/tactics/inv.ml b/tactics/inv.ml index ecb8eedac..632a29721 100644 --- a/tactics/inv.ml +++ b/tactics/inv.ml @@ -90,8 +90,7 @@ let make_inv_predicate env evd indf realargs id status concl = | None -> let sort = get_sort_family_of env !evd concl in let sort = Evarutil.evd_comb1 (Evd.fresh_sort_in_family env) evd sort in - let p = make_arity env true indf sort in - let p = EConstr.of_constr p in + let p = make_arity env !evd true indf sort in let evd',(p,ptyp) = Unification.abstract_list_all env !evd p concl (realargs@[mkVar id]) in evd := evd'; p in @@ -133,11 +132,7 @@ let make_inv_predicate env evd indf realargs id status concl = build_concl eqns args (succ n) restlist in let (newconcl, args) = build_concl [] [] 0 realargs in - let name_context env ctx = - let map f c = List.map (fun d -> Termops.map_rel_decl f d) c in - map EConstr.of_constr (name_context env (map EConstr.Unsafe.to_constr ctx)) - in - let predicate = it_mkLambda_or_LetIn newconcl (name_context env hyps) in + let predicate = it_mkLambda_or_LetIn newconcl (name_context env !evd hyps) in let _ = Evarutil.evd_comb1 (Typing.type_of env) evd predicate in (* OK - this predicate should now be usable by res_elimination_then to do elimination on the conclusion. *) diff --git a/tactics/leminv.ml b/tactics/leminv.ml index d7c396179..d864e547c 100644 --- a/tactics/leminv.ml +++ b/tactics/leminv.ml @@ -119,11 +119,11 @@ let max_prefix_sign lid sign = let rec add_prods_sign env sigma t = match EConstr.kind sigma (whd_all env sigma t) with | Prod (na,c1,b) -> - let id = id_of_name_using_hdchar env (EConstr.Unsafe.to_constr t) na in + let id = id_of_name_using_hdchar env sigma t na in let b'= subst1 (mkVar id) b in add_prods_sign (push_named (LocalAssum (id,c1)) env) sigma b' | LetIn (na,c1,t1,b) -> - let id = id_of_name_using_hdchar env (EConstr.Unsafe.to_constr t) na in + let id = id_of_name_using_hdchar env sigma t na in let b'= subst1 (mkVar id) b in add_prods_sign (push_named (LocalDef (id,c1,t1)) env) sigma b' | _ -> (env,t) @@ -147,8 +147,7 @@ let compute_first_inversion_scheme env sigma ind sort dep_option = let p = next_ident_away (Id.of_string "P") allvars in let pty,goal = if dep_option then - let pty = make_arity env true indf sort in - let pty = EConstr.of_constr pty in + let pty = make_arity env sigma true indf sort in let goal = mkProd (Anonymous, mkAppliedInd ind, applist(mkVar p,realargs@[mkRel 1])) diff --git a/tactics/tactics.ml b/tactics/tactics.ml index a29803251..13ffbc52f 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -412,16 +412,13 @@ let id_of_name_with_default id = function let default_id_of_sort s = if Sorts.is_small s then default_small_ident else default_type_ident -let id_of_name_using_hdchar env c name = - id_of_name_using_hdchar env (EConstr.Unsafe.to_constr c) name - let default_id env sigma decl = let open Context.Rel.Declaration in match decl with | LocalAssum (name,t) -> let dft = default_id_of_sort (Retyping.get_sort_of env sigma t) in id_of_name_with_default dft name - | LocalDef (name,b,_) -> id_of_name_using_hdchar env b name + | LocalDef (name,b,_) -> id_of_name_using_hdchar env sigma b name (* Non primitive introduction tactics are treated by intro_then_gen There is possibly renaming, with possibly names to avoid and @@ -1075,14 +1072,14 @@ let intros_replacing ids = (* User-level introduction tactics *) -let lookup_hypothesis_as_renamed env ccl = function - | AnonHyp n -> Detyping.lookup_index_as_renamed env (EConstr.Unsafe.to_constr ccl) n - | NamedHyp id -> Detyping.lookup_name_as_displayed env (EConstr.Unsafe.to_constr ccl) id +let lookup_hypothesis_as_renamed env sigma ccl = function + | AnonHyp n -> Detyping.lookup_index_as_renamed env sigma ccl n + | NamedHyp id -> Detyping.lookup_name_as_displayed env sigma ccl id let lookup_hypothesis_as_renamed_gen red h gl = let env = Proofview.Goal.env gl in let rec aux ccl = - match lookup_hypothesis_as_renamed env ccl h with + match lookup_hypothesis_as_renamed env (Tacmach.New.project gl) ccl h with | None when red -> let (redfun, _) = Redexpr.reduction_of_red_expr env (Red true) in let Sigma (c, _, _) = redfun.e_redfun env (Proofview.Goal.sigma gl) ccl in @@ -1350,7 +1347,7 @@ let enforce_prop_bound_names rename tac = if Retyping.get_sort_family_of env sigma t = InProp then (* "very_standard" says that we should have "H" names only, but this would break compatibility even more... *) - let s = match Namegen.head_name (EConstr.Unsafe.to_constr t) with + let s = match Namegen.head_name sigma t with | Some id when not very_standard -> string_of_id id | _ -> "" in Name (add_suffix Namegen.default_prop_ident s) @@ -2768,7 +2765,7 @@ let enough_by na t tac = forward false (Some (Some tac)) (ipat_of_name na) t (* Compute a name for a generalization *) -let generalized_name sigma c t ids cl = function +let generalized_name env sigma c t ids cl = function | Name id as na -> if Id.List.mem id ids then user_err (pr_id id ++ str " is already used."); @@ -2783,7 +2780,7 @@ let generalized_name sigma c t ids cl = function (* On ne s'etait pas casse la tete : on avait pris pour nom de variable la premiere lettre du type, meme si "c" avait ete une constante dont on aurait pu prendre directement le nom *) - named_hd (Global.env()) (EConstr.Unsafe.to_constr t) Anonymous + named_hd env sigma t Anonymous (* Abstract over [c] in [forall x1:A1(c)..xi:Ai(c).T(c)] producing [forall x, x1:A1(x1), .., xi:Ai(x). T(x)] with all [c] abtracted in [Ai] @@ -2795,7 +2792,7 @@ let generalize_goal_gen env sigma ids i ((occs,c,b),na) t cl = let dummy_prod = it_mkProd_or_LetIn mkProp decls in let newdecls,_ = decompose_prod_n_assum sigma i (subst_term_gen sigma EConstr.eq_constr_nounivs c dummy_prod) in let cl',sigma' = subst_closed_term_occ env sigma (AtOccs occs) c (it_mkProd_or_LetIn cl newdecls) in - let na = generalized_name sigma c t ids cl' na in + let na = generalized_name env sigma c t ids cl' na in let decl = match b with | None -> LocalAssum (na,t) | Some b -> LocalDef (na,b,t) @@ -3230,7 +3227,7 @@ let atomize_param_of_ind_then (indref,nparams,_) hyp0 tac = | Var id -> id | _ -> let type_of = Tacmach.New.pf_unsafe_type_of gl in - id_of_name_using_hdchar (Global.env()) (type_of c) Anonymous in + id_of_name_using_hdchar (Global.env()) sigma (type_of c) Anonymous in let x = fresh_id_in_env avoid id env in Tacticals.New.tclTHEN (letin_tac None (Name x) c None allHypsAndConcl) @@ -4434,7 +4431,7 @@ let induction_gen clear_flag isrec with_evars elim declaring the induction argument as a new local variable *) let id = (* Type not the right one if partially applied but anyway for internal use*) - let x = id_of_name_using_hdchar (Global.env()) t Anonymous in + let x = id_of_name_using_hdchar (Global.env()) evd t Anonymous in new_fresh_id [] x gl in let info_arg = (is_arg_pure_hyp, not enough_applied) in pose_induction_arg_then @@ -4471,7 +4468,7 @@ let induction_gen_l isrec with_evars elim names lc = let type_of = Tacmach.New.pf_unsafe_type_of gl in let sigma = Tacmach.New.project gl in let x = - id_of_name_using_hdchar (Global.env()) (type_of c) Anonymous in + id_of_name_using_hdchar (Global.env()) sigma (type_of c) Anonymous in let id = new_fresh_id [] x gl in let newl' = List.map (fun r -> replace_term sigma c (mkVar id) r) l' in diff --git a/test-suite/output/inference.out b/test-suite/output/inference.out index 576fbd7c0..c0eede99e 100644 --- a/test-suite/output/inference.out +++ b/test-suite/output/inference.out @@ -6,13 +6,13 @@ fun e : option L => match e with : option L -> option L fun (m n p : nat) (H : S m <= S n + p) => le_S_n m (n + p) H : forall m n p : nat, S m <= S n + p -> m <= n + p -fun n : nat => let x := A n : T n in ?y ?y0 : T n +fun n : nat => let x := A n : T n in ?t ?y : T n : forall n : nat, T n where -?y : [n : nat x := A n : T n |- ?T -> T n] -?y0 : [n : nat x := A n : T n |- ?T] -fun n : nat => ?y ?y0 : T n +?t : [n : nat x := A n : T n |- ?T -> T n] +?y : [n : nat x := A n : T n |- ?T] +fun n : nat => ?t ?y : T n : forall n : nat, T n where -?y : [n : nat |- ?T -> T n] -?y0 : [n : nat |- ?T] +?t : [n : nat |- ?T -> T n] +?y : [n : nat |- ?T] diff --git a/toplevel/himsg.ml b/toplevel/himsg.ml index 6982205fd..dea1f22d9 100644 --- a/toplevel/himsg.ml +++ b/toplevel/himsg.ml @@ -39,7 +39,7 @@ let contract env sigma lc = env | _ -> let t = Vars.substl !l (RelDecl.get_type decl) in - let decl = decl |> RelDecl.map_name (named_hd env (EConstr.Unsafe.to_constr t)) |> RelDecl.map_value (Vars.substl !l) |> RelDecl.set_type t in + let decl = decl |> RelDecl.map_name (named_hd env sigma t) |> RelDecl.map_value (Vars.substl !l) |> RelDecl.set_type t in l := (mkRel 1) :: List.map (Vars.lift 1) !l; push_rel decl env in @@ -206,7 +206,7 @@ let pr_puniverses f env (c,u) = let explain_elim_arity env sigma ind sorts c pj okinds = let open EConstr in - let env = make_all_name_different env in + let env = make_all_name_different env sigma in let pi = pr_inductive env (fst ind) in let pc = pr_leconstr_env env sigma c in let msg = match okinds with @@ -241,7 +241,7 @@ let explain_elim_arity env sigma ind sorts c pj okinds = let explain_case_not_inductive env sigma cj = let cj = Evarutil.j_nf_evar sigma cj in - let env = make_all_name_different env in + let env = make_all_name_different env sigma in let pc = pr_leconstr_env env sigma cj.uj_val in let pct = pr_leconstr_env env sigma cj.uj_type in match EConstr.kind sigma cj.uj_type with @@ -254,7 +254,7 @@ let explain_case_not_inductive env sigma cj = let explain_number_branches env sigma cj expn = let cj = Evarutil.j_nf_evar sigma cj in - let env = make_all_name_different env in + let env = make_all_name_different env sigma in let pc = pr_leconstr_env env sigma cj.uj_val in let pct = pr_leconstr_env env sigma cj.uj_type in str "Matching on term" ++ brk(1,1) ++ pc ++ spc () ++ @@ -263,7 +263,7 @@ let explain_number_branches env sigma cj expn = let explain_ill_formed_branch env sigma c ci actty expty = let simp t = Reductionops.nf_betaiota sigma (Evarutil.nf_evar sigma t) in - let env = make_all_name_different env in + let env = make_all_name_different env sigma in let pc = pr_leconstr_env env sigma c in let pa, pe = pr_explicit env sigma (simp actty) (simp expty) in strbrk "In pattern-matching on term" ++ brk(1,1) ++ pc ++ @@ -300,7 +300,7 @@ let explain_unification_error env sigma p1 p2 = function (* Error speaks from itself *) [] | ConversionFailed (env,t1,t2) -> if EConstr.eq_constr sigma t1 p1 && EConstr.eq_constr sigma t2 p2 then [] else - let env = make_all_name_different env in + let env = make_all_name_different env sigma in let t1 = Evarutil.nf_evar sigma t1 in let t2 = Evarutil.nf_evar sigma t2 in if not (EConstr.eq_constr sigma t1 p1) || not (EConstr.eq_constr sigma t2 p2) then @@ -328,7 +328,7 @@ let explain_unification_error env sigma p1 p2 = function let u = EConstr.of_constr u in let t = Evarutil.nf_evar sigma t in let u = Evarutil.nf_evar sigma u in - let env = make_all_name_different env in + let env = make_all_name_different env sigma in (strbrk "cannot satisfy constraint " ++ pr_leconstr_env env sigma t ++ str " == " ++ pr_leconstr_env env sigma u) :: aux t u e @@ -341,7 +341,7 @@ let explain_unification_error env sigma p1 p2 = function prlist_with_sep pr_semicolon (fun x -> x) l ++ str ")" let explain_actual_type env sigma j t reason = - let env = make_all_name_different env in + let env = make_all_name_different env sigma in let j = j_nf_betaiotaevar sigma j in let t = Reductionops.nf_betaiota sigma t in (** Actually print *) @@ -361,7 +361,7 @@ let explain_cant_apply_bad_type env sigma (n,exptyp,actualtyp) rator randl = let exptyp = Evarutil.nf_evar sigma exptyp in let actualtyp = Reductionops.nf_betaiota sigma actualtyp in let rator = Evarutil.j_nf_evar sigma rator in - let env = make_all_name_different env in + let env = make_all_name_different env sigma in let actualtyp, exptyp = pr_explicit env sigma actualtyp exptyp in let nargs = Array.length randl in (* let pe = pr_ne_context_of (str "in environment") env sigma in*) @@ -387,7 +387,7 @@ let explain_cant_apply_bad_type env sigma (n,exptyp,actualtyp) rator randl = let explain_cant_apply_not_functional env sigma rator randl = let randl = Evarutil.jv_nf_evar sigma randl in let rator = Evarutil.j_nf_evar sigma rator in - let env = make_all_name_different env in + let env = make_all_name_different env sigma in let nargs = Array.length randl in (* let pe = pr_ne_context_of (str "in environment") env sigma in*) let pr = pr_leconstr_env env sigma rator.uj_val in @@ -437,7 +437,7 @@ let explain_ill_formed_rec_body env sigma err names i fixenv vdefj = str "Recursive definition on" ++ spc () ++ pr_lconstr_env env sigma c ++ spc () ++ str "which should be an inductive type" | RecursionOnIllegalTerm(j,(arg_env, arg),le,lt) -> - let arg_env = make_all_name_different arg_env in + let arg_env = make_all_name_different arg_env sigma in let called = match names.(j) with Name id -> pr_id id @@ -503,7 +503,7 @@ let explain_ill_formed_rec_body env sigma err names i fixenv vdefj = pr_ne_context_of (str "In environment") env sigma ++ st ++ str "." ++ fnl () ++ (try (* May fail with unresolved globals. *) - let fixenv = make_all_name_different fixenv in + let fixenv = make_all_name_different fixenv sigma in let pvd = pr_lconstr_env fixenv sigma vdefj.(i).uj_val in str"Recursive definition is:" ++ spc () ++ pvd ++ str "." with e when CErrors.noncritical e -> mt ()) @@ -511,7 +511,7 @@ let explain_ill_formed_rec_body env sigma err names i fixenv vdefj = let explain_ill_typed_rec_body env sigma i names vdefj vargs = let vdefj = Evarutil.jv_nf_evar sigma vdefj in let vargs = Array.map (Evarutil.nf_evar sigma) vargs in - let env = make_all_name_different env in + let env = make_all_name_different env sigma in let pvd = pr_leconstr_env env sigma vdefj.(i).uj_val in let pvdt, pv = pr_explicit env sigma vdefj.(i).uj_type vargs.(i) in str "The " ++ @@ -521,13 +521,13 @@ let explain_ill_typed_rec_body env sigma i names vdefj vargs = str "while it should be" ++ spc () ++ pv ++ str "." let explain_cant_find_case_type env sigma c = - let env = make_all_name_different env in + let env = make_all_name_different env sigma in let pe = pr_leconstr_env env sigma c in str "Cannot infer the return type of pattern-matching on" ++ ws 1 ++ pe ++ str "." let explain_occur_check env sigma ev rhs = - let env = make_all_name_different env in + let env = make_all_name_different env sigma in let pt = pr_leconstr_env env sigma rhs in str "Cannot define " ++ pr_existential_key sigma ev ++ str " with term" ++ brk(1,1) ++ pt ++ spc () ++ str "that would depend on itself." @@ -624,7 +624,7 @@ let explain_wrong_case_info env (ind,u) ci = spc () ++ pc ++ str "." let explain_cannot_unify env sigma m n e = - let env = make_all_name_different env in + let env = make_all_name_different env sigma in let m = Evarutil.nf_evar sigma m in let n = Evarutil.nf_evar sigma n in let pm, pn = pr_explicit env sigma m n in @@ -695,7 +695,7 @@ let explain_unsatisfied_constraints env sigma cst = spc () ++ str "(maybe a bugged tactic)." let explain_type_error env sigma err = - let env = make_all_name_different env in + let env = make_all_name_different env sigma in match err with | UnboundRel n -> explain_unbound_rel env sigma n @@ -774,7 +774,7 @@ let pr_constraints printenv env sigma evars cstrs = (fun (ev, evi) -> fnl () ++ pr_existential_key sigma ev ++ str " : " ++ pr_lconstr_env env' sigma evi.evar_concl ++ fnl ()) l in - h 0 (pe ++ evs ++ pr_evar_constraints cstrs) + h 0 (pe ++ evs ++ pr_evar_constraints sigma cstrs) else let filter evk _ = Evar.Map.mem evk evars in pr_evar_map_filter ~with_univs:false filter sigma @@ -810,7 +810,7 @@ let explain_unsatisfiable_constraints env sigma constr comp = let explain_pretype_error env sigma err = let env = Evardefine.env_nf_betaiotaevar sigma env in - let env = make_all_name_different env in + let env = make_all_name_different env sigma in match err with | CantFindCaseType c -> explain_cant_find_case_type env sigma c | ActualTypeNotCoercible (j,t,e) -> @@ -1210,7 +1210,7 @@ let explain_recursion_scheme_error = function let explain_bad_pattern env sigma cstr ty = let ty = EConstr.to_constr sigma ty in - let env = make_all_name_different env in + let env = make_all_name_different env sigma in let pt = pr_lconstr_env env sigma ty in let pc = pr_constructor env cstr in str "Found the constructor " ++ pc ++ brk(1,1) ++ @@ -1255,7 +1255,7 @@ let explain_non_exhaustive env pats = let explain_cannot_infer_predicate env sigma typs = let inj c = EConstr.to_constr sigma c in let typs = Array.map_to_list (fun (c1, c2) -> (inj c1, inj c2)) typs in - let env = make_all_name_different env in + let env = make_all_name_different env sigma in let pr_branch (cstr,typ) = let cstr,_ = decompose_app cstr in str "For " ++ pr_lconstr_env env sigma cstr ++ str ": " ++ pr_lconstr_env env sigma typ diff --git a/toplevel/vernacentries.ml b/toplevel/vernacentries.ml index 55477cbcf..43ffa9ef3 100644 --- a/toplevel/vernacentries.ml +++ b/toplevel/vernacentries.ml @@ -1237,7 +1237,7 @@ let vernac_reserve bl = let env = Global.env() in let sigma = Evd.from_env env in let t,ctx = Constrintern.interp_type env sigma c in - let t = Detyping.detype false [] env (Evd.from_ctx ctx) t in + let t = Detyping.detype false [] env (Evd.from_ctx ctx) (EConstr.of_constr t) in let t,_ = Notation_ops.notation_constr_of_glob_constr (default_env ()) t in Reserve.declare_reserved_type idl t) in List.iter sb_decl bl -- cgit v1.2.3 From 4f66c854956bd05a24fd55c3d52fb669dbbb65e6 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Wed, 30 Nov 2016 15:06:42 +0100 Subject: Moving evar-normalization functions to EConstr. This removes code duplication between Evarutil and EConstr. --- engine/eConstr.ml | 37 +++++++++++++++++++------------------ engine/eConstr.mli | 3 +++ engine/evarutil.ml | 40 ++++------------------------------------ 3 files changed, 26 insertions(+), 54 deletions(-) diff --git a/engine/eConstr.ml b/engine/eConstr.ml index 7d4d17c63..2de6d7ae4 100644 --- a/engine/eConstr.ml +++ b/engine/eConstr.ml @@ -19,8 +19,10 @@ sig type t val kind : Evd.evar_map -> t -> (t, t) Constr.kind_of_term val kind_upto : Evd.evar_map -> constr -> (constr, types) Constr.kind_of_term +val whd_evar : Evd.evar_map -> t -> t val of_kind : (t, 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 end = @@ -32,40 +34,43 @@ let safe_evar_value sigma ev = try Some (Evd.existential_value sigma ev) with NotInstantiatedEvar | Not_found -> None -let rec kind sigma c = - let c0 = Constr.kind c in - match c0 with +let rec whd_evar sigma c = + match Constr.kind c with | Evar (evk, args) -> (match safe_evar_value sigma (evk, args) with - Some c -> kind sigma c - | None -> c0) + Some c -> whd_evar sigma c + | None -> c) | Sort (Type u) -> let u' = Evd.normalize_universe sigma u in - if u' == u then c0 else Sort (Type u') + 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 c0 else Const (c', u') + 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 c0 else Ind (i, u') + 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 c0 else Construct (co, u') - | App (c, args) when isEvar c -> + if u' == u then c else mkConstructU (co, u') + | App (f, args) when isEvar f -> (** Enforce smart constructor invariant on applications *) - let ev = destEvar c in + let ev = destEvar f in begin match safe_evar_value sigma ev with - | None -> c0 - | Some c -> kind sigma (mkApp (c, args)) + | None -> c + | Some f -> whd_evar sigma (mkApp (f, args)) end - | _ -> c0 + | _ -> c +let kind sigma c = Constr.kind (whd_evar sigma c) let kind_upto = kind 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 t = + Constr.map (fun t -> to_constr sigma t) (whd_evar sigma t) + end include API @@ -457,10 +462,6 @@ let fold sigma f acc c = match kind sigma c with | CoFix (_,(lna,tl,bl)) -> Array.fold_left2 (fun acc t b -> f (f acc t) b) acc tl bl -let to_constr sigma t = - let rec map c = Constr.map map (Constr.of_kind (kind_upto sigma c)) in - map (unsafe_to_constr t) - 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 diff --git a/engine/eConstr.mli b/engine/eConstr.mli index 83536d6f8..7992c0633 100644 --- a/engine/eConstr.mli +++ b/engine/eConstr.mli @@ -34,6 +34,8 @@ val kind : Evd.evar_map -> t -> (t, 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) 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}. *) @@ -146,6 +148,7 @@ 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} *) diff --git a/engine/evarutil.ml b/engine/evarutil.ml index ce6d06f23..1624dc93e 100644 --- a/engine/evarutil.ml +++ b/engine/evarutil.ml @@ -21,10 +21,6 @@ open Sigma.Notations module RelDecl = Context.Rel.Declaration module NamedDecl = Context.Named.Declaration -let safe_evar_info sigma evk = - try Some (Evd.find sigma evk) - with Not_found -> None - let safe_evar_value sigma ev = try Some (Evd.existential_value sigma ev) with NotInstantiatedEvar | Not_found -> None @@ -72,40 +68,12 @@ let rec flush_and_check_evars sigma c = let flush_and_check_evars sigma c = flush_and_check_evars sigma (EConstr.Unsafe.to_constr c) -(* let nf_evar_key = Profile.declare_profile "nf_evar" *) -(* let nf_evar = Profile.profile2 nf_evar_key Reductionops.nf_evar *) - -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 - (** Term exploration up to instantiation. *) -let kind_of_term_upto sigma t = - Constr.kind (whd_evar sigma t) +let kind_of_term_upto = EConstr.kind_upto -let rec nf_evar0 sigma t = Constr.map (fun t -> nf_evar0 sigma t) (whd_evar sigma t) -let whd_evar sigma c = EConstr.of_constr (whd_evar sigma (EConstr.Unsafe.to_constr c)) -let nf_evar sigma c = EConstr.of_constr (nf_evar0 sigma (EConstr.Unsafe.to_constr c)) +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; -- cgit v1.2.3 From d629ec7cd920b19a063b7198d4e5b92d91a5656b Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sat, 17 Dec 2016 18:14:45 +0100 Subject: Putting back the occur_meta_or_undefined_evar function in the old term API. This is another perfomance-critical function in unification. Putting it in the EConstr API was changing the heuristic, so better revert on that change. --- pretyping/unification.ml | 13 ++++++++++--- 1 file changed, 10 insertions(+), 3 deletions(-) diff --git a/pretyping/unification.ml b/pretyping/unification.ml index baa12db08..318a0b2cd 100644 --- a/pretyping/unification.ml +++ b/pretyping/unification.ml @@ -65,10 +65,17 @@ let _ = Goptions.declare_bool_option { } let occur_meta_or_undefined_evar evd c = - let rec occrec c = match EConstr.kind evd c with + (** This is performance-critical. Using the evar-insensitive API changes the + resulting heuristic. *) + let c = EConstr.Unsafe.to_constr c in + let rec occrec c = match kind_of_term c with | Meta _ -> raise Occur - | Evar _ -> raise Occur - | _ -> EConstr.iter evd occrec c + | Evar (ev,args) -> + (match evar_body (Evd.find evd ev) with + | Evar_defined c -> + occrec c; Array.iter occrec args + | Evar_empty -> raise Occur) + | _ -> Constr.iter occrec c in try occrec c; false with Occur | Not_found -> true let occur_meta_evd sigma mv c = -- cgit v1.2.3 From 63ae87d51456add79652b42b972d6be93b6119bc Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sat, 17 Dec 2016 15:50:10 +0100 Subject: Fix a mishandled exception in Omega. Due to the introduction of the monadic layer, an exception was raised at a later time and not caught properly. --- plugins/omega/coq_omega.ml | 4 +++- 1 file changed, 3 insertions(+), 1 deletion(-) diff --git a/plugins/omega/coq_omega.ml b/plugins/omega/coq_omega.ml index 8c2f0f53f..72aeb9066 100644 --- a/plugins/omega/coq_omega.ml +++ b/plugins/omega/coq_omega.ml @@ -1859,7 +1859,9 @@ let destructure_goal = let decidability = Tacmach.New.of_old decidability gl in let rec loop t = Proofview.tclEVARMAP >>= fun sigma -> - match destructurate_prop sigma t with + let prop () = Proofview.tclUNIT (destructurate_prop sigma t) in + Proofview.V82.wrap_exceptions prop >>= fun prop -> + match prop with | Kapp(Not,[t]) -> (Tacticals.New.tclTHEN (Tacticals.New.tclTHEN (unfold sp_not) intro) -- cgit v1.2.3 From 3c1cd2338fcddc4a6c0e97b0af53eb2b2f238c4a Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Thu, 15 Dec 2016 10:45:19 +0100 Subject: Removing most nf_enter in tactics. Now they are useless because all of the primitives are (should?) be evar-insensitive. --- ltac/extratactics.ml4 | 24 ++++---- ltac/g_class.ml4 | 6 +- ltac/rewrite.ml | 8 +-- ltac/tacinterp.ml | 30 +++++----- plugins/cc/cctac.ml | 44 +++++++------- plugins/omega/coq_omega.ml | 35 +++++++----- proofs/clenv.ml | 17 ++++-- proofs/clenv.mli | 11 ++-- proofs/clenvtac.ml | 4 +- tactics/auto.ml | 16 +++--- tactics/autorewrite.ml | 2 +- tactics/class_tactics.ml | 42 +++++++------- tactics/class_tactics.mli | 2 +- tactics/contradiction.ml | 2 +- tactics/eauto.ml | 17 +++--- tactics/elim.ml | 4 +- tactics/eqdecide.ml | 4 +- tactics/equality.ml | 24 ++++---- tactics/inv.ml | 14 ++--- tactics/leminv.ml | 5 +- tactics/tacticals.ml | 140 ++++++++++++++++++++++----------------------- tactics/tacticals.mli | 5 +- tactics/tactics.ml | 107 ++++++++++++++++++---------------- tactics/tactics.mli | 4 +- toplevel/auto_ind_decl.ml | 32 +++++------ 25 files changed, 305 insertions(+), 294 deletions(-) diff --git a/ltac/extratactics.ml4 b/ltac/extratactics.ml4 index 58c5823c5..e9d91fe47 100644 --- a/ltac/extratactics.ml4 +++ b/ltac/extratactics.ml4 @@ -354,7 +354,7 @@ let constr_flags = { Pretyping.expand_evars = true } let refine_tac ist simple c = - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> let concl = Proofview.Goal.concl gl in let env = Proofview.Goal.env gl in let flags = constr_flags in @@ -645,7 +645,7 @@ let subst_hole_with_term occ tc t = open Tacmach let hResolve id c occ t = - Proofview.Goal.nf_s_enter { s_enter = begin fun gl -> + Proofview.Goal.s_enter { s_enter = begin fun gl -> let sigma = Proofview.Goal.sigma gl in let sigma = Sigma.to_evar_map sigma in let env = Termops.clear_named_body id (Proofview.Goal.env gl) in @@ -708,7 +708,7 @@ END exception Found of unit Proofview.tactic let rewrite_except h = - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> let hyps = Tacmach.New.pf_ids_of_hyps gl in Tacticals.New.tclMAP (fun id -> if Id.equal id h then Proofview.tclUNIT () else Tacticals.New.tclTRY (Equality.general_rewrite_in true Locus.AllOccurrences true true id (mkVar h) false)) @@ -727,11 +727,11 @@ let refl_equal = should be replaced by a call to the tactic but I don't know how to call it before it is defined. *) let mkCaseEq a : unit Proofview.tactic = - Proofview.Goal.nf_enter { enter = begin fun gl -> - let type_of_a = Tacmach.New.of_old (fun g -> Tacmach.pf_unsafe_type_of g a) gl in + Proofview.Goal.enter { enter = begin fun gl -> + let type_of_a = Tacmach.New.pf_unsafe_type_of gl a in Tacticals.New.tclTHENLIST [Tactics.generalize [(mkApp(EConstr.of_constr (delayed_force refl_equal), [| type_of_a; a|]))]; - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> let concl = Proofview.Goal.concl gl in let env = Proofview.Goal.env gl in (** FIXME: this looks really wrong. Does anybody really use this tactic? *) @@ -743,16 +743,16 @@ let mkCaseEq a : unit Proofview.tactic = let case_eq_intros_rewrite x = - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> let n = nb_prod (Tacmach.New.project gl) (Proofview.Goal.concl gl) in (* Pp.msgnl (Printer.pr_lconstr x); *) Tacticals.New.tclTHENLIST [ mkCaseEq x; - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> let concl = Proofview.Goal.concl gl in let hyps = Tacmach.New.pf_ids_of_hyps gl in let n' = nb_prod (Tacmach.New.project gl) concl in - let h = Tacmach.New.of_old (fun g -> fresh_id hyps (Id.of_string "heq") g) gl in + let h = fresh_id_in_env hyps (Id.of_string "heq") (Proofview.Goal.env gl) in Tacticals.New.tclTHENLIST [ Tacticals.New.tclDO (n'-n-1) intro; introduction h; @@ -783,15 +783,15 @@ let destauto t = with Found tac -> tac let destauto_in id = - Proofview.Goal.nf_enter { enter = begin fun gl -> - let ctype = Tacmach.New.of_old (fun g -> Tacmach.pf_unsafe_type_of g (mkVar id)) gl in + Proofview.Goal.enter { enter = begin fun gl -> + let ctype = Tacmach.New.pf_unsafe_type_of gl (mkVar id) in (* Pp.msgnl (Printer.pr_lconstr (mkVar id)); *) (* Pp.msgnl (Printer.pr_lconstr (ctype)); *) destauto ctype end } TACTIC EXTEND destauto -| [ "destauto" ] -> [ Proofview.Goal.nf_enter { enter = begin fun gl -> destauto (Proofview.Goal.concl gl) end } ] +| [ "destauto" ] -> [ Proofview.Goal.enter { enter = begin fun gl -> destauto (Proofview.Goal.concl gl) end } ] | [ "destauto" "in" hyp(id) ] -> [ destauto_in id ] END diff --git a/ltac/g_class.ml4 b/ltac/g_class.ml4 index b7f0881f1..f47c03dfc 100644 --- a/ltac/g_class.ml4 +++ b/ltac/g_class.ml4 @@ -74,7 +74,7 @@ TACTIC EXTEND is_ground END TACTIC EXTEND autoapply - [ "autoapply" constr(c) "using" preident(i) ] -> [ Proofview.V82.tactic (autoapply c i) ] + [ "autoapply" constr(c) "using" preident(i) ] -> [ autoapply c i ] END (** TODO: DEPRECATE *) @@ -90,10 +90,10 @@ let rec eq_constr_mod_evars sigma x y = | _, _ -> compare_constr sigma (fun x y -> eq_constr_mod_evars sigma x y) x y let progress_evars t = - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> let concl = Proofview.Goal.concl gl in let check = - Proofview.Goal.nf_enter { enter = begin fun gl' -> + Proofview.Goal.enter { enter = begin fun gl' -> let sigma = Tacmach.New.project gl' in let newconcl = Proofview.Goal.concl gl' in if eq_constr_mod_evars sigma concl newconcl diff --git a/ltac/rewrite.ml b/ltac/rewrite.ml index bc57fcf56..810f02c4f 100644 --- a/ltac/rewrite.ml +++ b/ltac/rewrite.ml @@ -1543,7 +1543,7 @@ let rec insert_dependent env sigma decl accu hyps = match hyps with insert_dependent env sigma decl (ndecl :: accu) rem let assert_replacing id newt tac = - let prf = Proofview.Goal.nf_enter { enter = begin fun gl -> + let prf = Proofview.Goal.enter { enter = begin fun gl -> let concl = Proofview.Goal.concl gl in let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in @@ -1611,7 +1611,7 @@ let cl_rewrite_clause_newtac ?abs ?origsigma ~progress strat clause = Proofview.Unsafe.tclEVARS undef <*> convert_concl_no_check newt DEFAULTcast in - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> let concl = Proofview.Goal.concl gl in let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in @@ -2095,7 +2095,7 @@ let general_rewrite_flags = { under_lambdas = false; on_morphisms = true } (** Setoid rewriting when called with "rewrite" *) let general_s_rewrite cl l2r occs (c,l) ~new_goals = - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> let abs, evd, res, sort = get_hyp gl (c,l) cl l2r in let unify env evars t = unify_abs res l2r sort env evars t in let app = apply_rule unify occs in @@ -2129,7 +2129,7 @@ let not_declared env sigma ty rel = str ty ++ str" relation. Maybe you need to require the Coq.Classes.RelationClasses library") let setoid_proof ty fn fallback = - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in let concl = Proofview.Goal.concl gl in diff --git a/ltac/tacinterp.ml b/ltac/tacinterp.ml index 62d0cc529..167501de2 100644 --- a/ltac/tacinterp.ml +++ b/ltac/tacinterp.ml @@ -844,10 +844,10 @@ let rec message_of_value v = Ftactic.return (str "") else if has_type v (topwit wit_constr) then let v = out_gen (topwit wit_constr) v in - Ftactic.nf_enter {enter = begin fun gl -> Ftactic.return (pr_econstr_env (pf_env gl) (project gl) v) end } + Ftactic.enter {enter = begin fun gl -> Ftactic.return (pr_econstr_env (pf_env gl) (project gl) v) end } else if has_type v (topwit wit_constr_under_binders) then let c = out_gen (topwit wit_constr_under_binders) v in - Ftactic.nf_enter { enter = begin fun gl -> + Ftactic.enter { enter = begin fun gl -> Ftactic.return (pr_constr_under_binders_env (pf_env gl) (project gl) c) end } else if has_type v (topwit wit_unit) then @@ -860,21 +860,21 @@ let rec message_of_value v = let (c, sigma) = Tactics.run_delayed env sigma c in pr_econstr_env env sigma c in - Ftactic.nf_enter { enter = begin fun gl -> + Ftactic.enter { enter = begin fun gl -> Ftactic.return (Miscprint.pr_intro_pattern (fun c -> print (pf_env gl) (project gl) c) p) end } else if has_type v (topwit wit_constr_context) then let c = out_gen (topwit wit_constr_context) v in - Ftactic.nf_enter { enter = begin fun gl -> Ftactic.return (pr_econstr_env (pf_env gl) (project gl) c) end } + Ftactic.enter { enter = begin fun gl -> Ftactic.return (pr_econstr_env (pf_env gl) (project gl) c) end } else if has_type v (topwit wit_uconstr) then let c = out_gen (topwit wit_uconstr) v in - Ftactic.nf_enter { enter = begin fun gl -> + Ftactic.enter { enter = begin fun gl -> Ftactic.return (pr_closed_glob_env (pf_env gl) (project gl) c) end } else if has_type v (topwit wit_var) then let id = out_gen (topwit wit_var) v in - Ftactic.nf_enter { enter = begin fun gl -> Ftactic.return (pr_id id) end } + Ftactic.enter { enter = begin fun gl -> Ftactic.return (pr_id id) end } else match Value.to_list v with | Some l -> Ftactic.List.map message_of_value l >>= fun l -> @@ -1213,7 +1213,7 @@ and eval_tactic ist tac : unit Proofview.tactic = match tac with tclSHOWHYPS (Proofview.V82.of_tactic (interp_tactic ist tac)) end | TacAbstract (tac,ido) -> - Proofview.Goal.nf_enter { enter = begin fun gl -> Tactics.tclABSTRACT + Proofview.Goal.enter { enter = begin fun gl -> Tactics.tclABSTRACT (Option.map (interp_ident ist (pf_env gl) (project gl)) ido) (interp_tactic ist tac) end } | TacThen (t1,t) -> @@ -1521,7 +1521,7 @@ and interp_match ist lz constr lmr = (* Interprets the Match Context expressions *) and interp_match_goal ist lz lr lmr = - Ftactic.nf_enter { enter = begin fun gl -> + Ftactic.enter { enter = begin fun gl -> let sigma = project gl in let env = Proofview.Goal.env gl in let hyps = Proofview.Goal.hyps gl in @@ -1750,8 +1750,7 @@ and interp_atomic ist tac : unit Proofview.tactic = (Tactics.generalize_gen cl)) sigma end } | TacLetTac (na,c,clp,b,eqpat) -> - Proofview.V82.nf_evar_goals <*> - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> let env = Proofview.Goal.env gl in let sigma = project gl in let clp = interp_clause ist env sigma clp in @@ -1788,7 +1787,6 @@ and interp_atomic ist tac : unit Proofview.tactic = | TacInductionDestruct (isrec,ev,(l,el)) -> (* spiwack: some unknown part of destruct needs the goal to be prenormalised. *) - Proofview.V82.nf_evar_goals <*> Proofview.Goal.nf_s_enter { s_enter = begin fun gl -> let env = Proofview.Goal.env gl in let sigma = project gl in @@ -1824,8 +1822,7 @@ and interp_atomic ist tac : unit Proofview.tactic = | TacChange (None,c,cl) -> (* spiwack: until the tactic is in the monad *) Proofview.Trace.name_tactic (fun () -> Pp.str"") begin - Proofview.V82.nf_evar_goals <*> - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> let is_onhyps = match cl.onhyps with | None | Some [] -> true | _ -> false @@ -1854,7 +1851,6 @@ and interp_atomic ist tac : unit Proofview.tactic = | TacChange (Some op,c,cl) -> (* spiwack: until the tactic is in the monad *) Proofview.Trace.name_tactic (fun () -> Pp.str"") begin - Proofview.V82.nf_evar_goals <*> Proofview.Goal.enter { enter = begin fun gl -> let env = Proofview.Goal.env gl in let sigma = project gl in @@ -1899,7 +1895,7 @@ and interp_atomic ist tac : unit Proofview.tactic = by)) end } | TacInversion (DepInversion (k,c,ids),hyp) -> - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> let env = Proofview.Goal.env gl in let sigma = project gl in let (sigma,c_interp) = @@ -2050,7 +2046,7 @@ let interp_constr_with_bindings' ist c = Ftactic.return { delayed = fun env sigm Sigma.Unsafe.of_pair (c, sigma) } -let interp_destruction_arg' ist c = Ftactic.nf_enter { enter = begin fun gl -> +let interp_destruction_arg' ist c = Ftactic.enter { enter = begin fun gl -> Ftactic.return (interp_destruction_arg ist gl c) end } @@ -2080,7 +2076,7 @@ let () = register_interp0 wit_ltac interp let () = - register_interp0 wit_uconstr (fun ist c -> Ftactic.nf_enter { enter = begin fun gl -> + register_interp0 wit_uconstr (fun ist c -> Ftactic.enter { enter = begin fun gl -> Ftactic.return (interp_uconstr ist (Proofview.Goal.env gl) (Tacmach.New.project gl) c) end }) diff --git a/plugins/cc/cctac.ml b/plugins/cc/cctac.ml index 5d894c677..c9c904e35 100644 --- a/plugins/cc/cctac.ml +++ b/plugins/cc/cctac.ml @@ -182,9 +182,10 @@ let litteral_of_constr env sigma term= (* store all equalities from the context *) let make_prb gls depth additionnal_terms = + let open Tacmach.New in let env=pf_env gls in - let sigma=sig_sig gls in - let state = empty depth gls in + let sigma=project gls in + let state = empty depth {it = Proofview.Goal.goal (Proofview.Goal.assume gls); sigma } in let pos_hyps = ref [] in let neg_hyps =ref [] in List.iter @@ -196,7 +197,7 @@ let make_prb gls depth additionnal_terms = let id = NamedDecl.get_id decl in begin let cid=Constr.mkVar id in - match litteral_of_constr env sigma (EConstr.of_constr (NamedDecl.get_type decl)) with + match litteral_of_constr env sigma (NamedDecl.get_type decl) with `Eq (t,a,b) -> add_equality state cid a b | `Neq (t,a,b) -> add_disequality state (Hyp cid) a b | `Other ph -> @@ -213,7 +214,7 @@ let make_prb gls depth additionnal_terms = neg_hyps:=(cid,nh):: !neg_hyps | `Rule patts -> add_quant state id true patts | `Nrule patts -> add_quant state id false patts - end) (Environ.named_context_of_val (Goal.V82.nf_hyps gls.sigma gls.it)); + end) (Proofview.Goal.hyps gls); begin match atom_of_constr env sigma (pf_concl gls) with `Eq (t,a,b) -> add_disequality state Goal a b @@ -227,6 +228,7 @@ let make_prb gls depth additionnal_terms = (* indhyps builds the array of arrays of constructor hyps for (ind largs) *) let build_projection intype (cstr:pconstructor) special default gls= + let open Tacmach.New in let ci= (snd(fst cstr)) in let body=Equality.build_selector (pf_env gls) (project gls) ci (mkRel 1) intype special default in let id=pf_get_new_id (Id.of_string "t") gls in @@ -266,7 +268,7 @@ let refresh_universes ty k = let constr_of_term c = EConstr.of_constr (constr_of_term c) let rec proof_tac p : unit Proofview.tactic = - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> let type_of t = Tacmach.New.pf_unsafe_type_of gl t in try (* type_of can raise exceptions *) match p.p_rule with @@ -296,7 +298,7 @@ let rec proof_tac p : unit Proofview.tactic = refresh_universes (type_of tf1) (fun typf -> refresh_universes (type_of tx1) (fun typx -> refresh_universes (type_of (mkApp (tf1,[|tx1|]))) (fun typfx -> - let id = Tacmach.New.of_old (fun gls -> pf_get_new_id (Id.of_string "f") gls) gl in + let id = Tacmach.New.pf_get_new_id (Id.of_string "f") gl in let appx1 = mkLambda(Name id,typf,mkApp(mkRel 1,[|tx1|])) in let lemma1 = app_global _f_equal [|typf;typfx;appx1;tf1;tf2;_M 1|] in let lemma2 = app_global _f_equal [|typx;typfx;tf2;tx1;tx2;_M 1|] in @@ -322,7 +324,7 @@ let rec proof_tac p : unit Proofview.tactic = refresh_universes (type_of ti) (fun intype -> refresh_universes (type_of default) (fun outtype -> let proj = - Tacmach.New.of_old (build_projection intype cstr special default) gl + build_projection intype cstr special default gl in let injt= app_global _f_equal [|intype;outtype;proj;ti;tj;_M 1|] in @@ -331,9 +333,9 @@ let rec proof_tac p : unit Proofview.tactic = end } let refute_tac c t1 t2 p = - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> let tt1=constr_of_term t1 and tt2=constr_of_term t2 in - let hid = Tacmach.New.of_old (pf_get_new_id (Id.of_string "Heq")) gl in + let hid = Tacmach.New.pf_get_new_id (Id.of_string "Heq") gl in let false_t=mkApp (c,[|mkVar hid|]) in let k intype = let neweq= new_app_global _eq [|intype;tt1;tt2|] in @@ -347,12 +349,12 @@ let refine_exact_check c gl = Tacticals.tclTHEN (Refiner.tclEVARS evm) (Proofview.V82.of_tactic (exact_check c)) gl let convert_to_goal_tac c t1 t2 p = - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> let tt1=constr_of_term t1 and tt2=constr_of_term t2 in let k sort = let neweq= new_app_global _eq [|sort;tt1;tt2|] in - let e = Tacmach.New.of_old (pf_get_new_id (Id.of_string "e")) gl in - let x = Tacmach.New.of_old (pf_get_new_id (Id.of_string "X")) gl in + let e = Tacmach.New.pf_get_new_id (Id.of_string "e") gl in + let x = Tacmach.New.pf_get_new_id (Id.of_string "X") gl in let identity=mkLambda (Name x,sort,mkRel 1) in let endt=app_global _eq_rect [|sort;tt1;identity;c;tt2;mkVar e|] in Tacticals.New.tclTHENS (neweq (assert_before (Name e))) @@ -361,9 +363,9 @@ let convert_to_goal_tac c t1 t2 p = end } let convert_to_hyp_tac c1 t1 c2 t2 p = - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> let tt2=constr_of_term t2 in - let h = Tacmach.New.of_old (pf_get_new_id (Id.of_string "H")) gl in + let h = Tacmach.New.pf_get_new_id (Id.of_string "H") gl in let false_t=mkApp (c2,[|mkVar h|]) in Tacticals.New.tclTHENS (assert_before (Name h) tt2) [convert_to_goal_tac c1 t1 t2 p; @@ -371,11 +373,11 @@ let convert_to_hyp_tac c1 t1 c2 t2 p = end } let discriminate_tac (cstr,u as cstru) p = - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> let t1=constr_of_term p.p_lhs and t2=constr_of_term p.p_rhs in let env = Proofview.Goal.env gl in let concl = Proofview.Goal.concl gl in - let xid = Tacmach.New.of_old (pf_get_new_id (Id.of_string "X")) gl in + let xid = Tacmach.New.pf_get_new_id (Id.of_string "X") gl in let identity = Universes.constr_of_global (Lazy.force _I) in let identity = EConstr.of_constr identity in let trivial = Universes.constr_of_global (Lazy.force _True) in @@ -385,8 +387,8 @@ let discriminate_tac (cstr,u as cstru) p = let evm, outtype = Evd.new_sort_variable Evd.univ_flexible evm in let outtype = mkSort outtype in let pred = mkLambda(Name xid,outtype,mkRel 1) in - let hid = Tacmach.New.of_old (pf_get_new_id (Id.of_string "Heq")) gl in - let proj = Tacmach.New.of_old (build_projection intype cstru trivial concl) gl in + let hid = Tacmach.New.pf_get_new_id (Id.of_string "Heq") gl in + let proj = build_projection intype cstru trivial concl gl in let injt=app_global _f_equal [|intype;outtype;proj;t1;t2;mkVar hid|] in let endt k = @@ -409,11 +411,11 @@ let build_term_to_complete uf meta pac = applist (mkConstructU cinfo.ci_constr, all_args) let cc_tactic depth additionnal_terms = - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> let sigma = Tacmach.New.project gl in Coqlib.check_required_library Coqlib.logic_module_name; let _ = debug (fun () -> Pp.str "Reading subgoal ...") in - let state = Tacmach.New.of_old (fun gls -> make_prb gls depth additionnal_terms) gl in + let state = make_prb gl depth additionnal_terms in let _ = debug (fun () -> Pp.str "Problem built, solving ...") in let sol = execute true state in let _ = debug (fun () -> Pp.str "Computation completed.") in @@ -498,7 +500,7 @@ let mk_eq f c1 c2 k = end }) let f_equal = - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> let concl = Proofview.Goal.concl gl in let sigma = Tacmach.New.project gl in let cut_eq c1 c2 = diff --git a/plugins/omega/coq_omega.ml b/plugins/omega/coq_omega.ml index 72aeb9066..adf926958 100644 --- a/plugins/omega/coq_omega.ml +++ b/plugins/omega/coq_omega.ml @@ -38,7 +38,7 @@ open OmegaSolver (* Added by JCF, 09/03/98 *) let elim_id id = - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> simplest_elim (Tacmach.New.pf_global id gl) end } let resolve_id id gl = Proofview.V82.of_tactic (apply (pf_global gl id)) gl @@ -1391,7 +1391,10 @@ let normalize_equation sigma id flag theorem pos t t1 t2 (tactic,defs) = else (tactic,defs) +let pf_nf gl c = Tacmach.New.pf_apply Tacred.simpl gl c + let destructure_omega gl tac_def (id,c) = + let open Tacmach.New in let sigma = project gl in if String.equal (atompart_of_id id) "State" then tac_def @@ -1433,10 +1436,10 @@ let reintroduce id = open Proofview.Notations let coq_omega = - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> clear_constr_tables (); let hyps_types = Tacmach.New.pf_hyps_types gl in - let destructure_omega = Tacmach.New.of_old destructure_omega gl in + let destructure_omega = destructure_omega gl in let tactic_normalisation, system = List.fold_left destructure_omega ([],[]) hyps_types in let prelude,sys = @@ -1486,7 +1489,7 @@ let coq_omega = let coq_omega = coq_omega let nat_inject = - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> let is_conv = Tacmach.New.pf_apply Reductionops.is_conv gl in let rec explore p t : unit Proofview.tactic = Proofview.tclEVARMAP >>= fun sigma -> @@ -1657,6 +1660,7 @@ let not_binop = function exception Undecidable let rec decidability gl t = + let open Tacmach.New in match destructurate_prop (project gl) t with | Kapp(Or,[t1;t2]) -> mkApp (Lazy.force coq_dec_or, [| t1; t2; @@ -1687,22 +1691,25 @@ let rec decidability gl t = | Kapp(True,[]) -> Lazy.force coq_dec_True | _ -> raise Undecidable +let fresh_id avoid id gl = + fresh_id_in_env avoid id (Proofview.Goal.env gl) + let onClearedName id tac = (* We cannot ensure that hyps can be cleared (because of dependencies), *) (* so renaming may be necessary *) Tacticals.New.tclTHEN (Tacticals.New.tclTRY (clear [id])) - (Proofview.Goal.nf_enter { enter = begin fun gl -> - let id = Tacmach.New.of_old (fresh_id [] id) gl in + (Proofview.Goal.enter { enter = begin fun gl -> + let id = fresh_id [] id gl in Tacticals.New.tclTHEN (introduction id) (tac id) end }) let onClearedName2 id tac = Tacticals.New.tclTHEN (Tacticals.New.tclTRY (clear [id])) - (Proofview.Goal.nf_enter { enter = begin fun gl -> - let id1 = Tacmach.New.of_old (fresh_id [] (add_suffix id "_left")) gl in - let id2 = Tacmach.New.of_old (fresh_id [] (add_suffix id "_right")) gl in + (Proofview.Goal.enter { enter = begin fun gl -> + let id1 = fresh_id [] (add_suffix id "_left") gl in + let id2 = fresh_id [] (add_suffix id "_right") gl in Tacticals.New.tclTHENLIST [ introduction id1; introduction id2; tac id1 id2 ] end }) @@ -1712,10 +1719,10 @@ let rec is_Prop sigma c = match EConstr.kind sigma c with | _ -> false let destructure_hyps = - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> let type_of = Tacmach.New.pf_unsafe_type_of gl in - let decidability = Tacmach.New.of_old decidability gl in - let pf_nf = Tacmach.New.of_old pf_nf gl in + let decidability = decidability gl in + let pf_nf = pf_nf gl in let rec loop = function | [] -> (Tacticals.New.tclTHEN nat_inject coq_omega) | decl::lit -> @@ -1854,9 +1861,9 @@ let destructure_hyps = end } let destructure_goal = - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> let concl = Proofview.Goal.concl gl in - let decidability = Tacmach.New.of_old decidability gl in + let decidability = decidability gl in let rec loop t = Proofview.tclEVARMAP >>= fun sigma -> let prop () = Proofview.tclUNIT (destructurate_prop sigma t) in diff --git a/proofs/clenv.ml b/proofs/clenv.ml index 5645850b2..f9ebc4233 100644 --- a/proofs/clenv.ml +++ b/proofs/clenv.ml @@ -129,11 +129,13 @@ let mk_clenv_from_env env sigma n (c,cty) = env = env } let mk_clenv_from_n gls n (c,cty) = - mk_clenv_from_env (pf_env gls) gls.sigma n (c, cty) + let env = Proofview.Goal.env gls in + let sigma = Tacmach.New.project gls in + mk_clenv_from_env env sigma n (c, cty) let mk_clenv_from gls = mk_clenv_from_n gls None -let mk_clenv_type_of gls t = mk_clenv_from gls (t,pf_type_of gls t) +let mk_clenv_type_of gls t = mk_clenv_from gls (t,Tacmach.New.pf_unsafe_type_of gls t) (******************************************************************) @@ -263,8 +265,7 @@ let clenv_unify ?(flags=default_unify_flags ()) cv_pb t1 t2 clenv = let clenv_unify_meta_types ?(flags=default_unify_flags ()) clenv = { clenv with evd = w_unify_meta_types ~flags:flags clenv.env clenv.evd } -let clenv_unique_resolver ?(flags=default_unify_flags ()) clenv gl = - let concl = Goal.V82.concl clenv.evd (sig_it gl) in +let clenv_unique_resolver_gen ?(flags=default_unify_flags ()) clenv concl = if isMeta clenv.evd (fst (decompose_app_vect clenv.evd (whd_nored clenv.evd clenv.templtyp.rebus))) then clenv_unify CUMUL ~flags (clenv_type clenv) concl (clenv_unify_meta_types ~flags clenv) @@ -272,6 +273,14 @@ let clenv_unique_resolver ?(flags=default_unify_flags ()) clenv gl = clenv_unify CUMUL ~flags (meta_reducible_instance clenv.evd clenv.templtyp) concl clenv +let old_clenv_unique_resolver ?flags clenv gl = + let concl = Goal.V82.concl clenv.evd (sig_it gl) in + clenv_unique_resolver_gen ?flags clenv concl + +let clenv_unique_resolver ?flags clenv gl = + let concl = Proofview.Goal.concl gl in + clenv_unique_resolver_gen ?flags clenv concl + let adjust_meta_source evd mv = function | loc,Evar_kinds.VarInstance id -> let rec match_name c l = diff --git a/proofs/clenv.mli b/proofs/clenv.mli index f7ff4bed3..4bcd50591 100644 --- a/proofs/clenv.mli +++ b/proofs/clenv.mli @@ -41,10 +41,10 @@ val clenv_nf_meta : clausenv -> EConstr.constr -> EConstr.constr (** type of a meta in clenv context *) val clenv_meta_type : clausenv -> metavariable -> types -val mk_clenv_from : Goal.goal sigma -> EConstr.constr * EConstr.types -> clausenv +val mk_clenv_from : ('a, 'r) Proofview.Goal.t -> EConstr.constr * EConstr.types -> clausenv val mk_clenv_from_n : - Goal.goal sigma -> int option -> EConstr.constr * EConstr.types -> clausenv -val mk_clenv_type_of : Goal.goal sigma -> EConstr.constr -> clausenv + ('a, 'r) Proofview.Goal.t -> int option -> EConstr.constr * EConstr.types -> clausenv +val mk_clenv_type_of : ('a, 'r) Proofview.Goal.t -> EConstr.constr -> clausenv val mk_clenv_from_env : env -> evar_map -> int option -> EConstr.constr * EConstr.types -> clausenv (** Refresh the universes in a clenv *) @@ -62,9 +62,12 @@ val clenv_unify : ?flags:unify_flags -> conv_pb -> constr -> constr -> clausenv -> clausenv (** unifies the concl of the goal with the type of the clenv *) -val clenv_unique_resolver : +val old_clenv_unique_resolver : ?flags:unify_flags -> clausenv -> Goal.goal sigma -> clausenv +val clenv_unique_resolver : + ?flags:unify_flags -> clausenv -> ('a, 'r) Proofview.Goal.t -> clausenv + val clenv_dependent : clausenv -> metavariable list val clenv_pose_metas_as_evars : clausenv -> metavariable list -> clausenv diff --git a/proofs/clenvtac.ml b/proofs/clenvtac.ml index 32c887ff5..0722ea047 100644 --- a/proofs/clenvtac.ml +++ b/proofs/clenvtac.ml @@ -105,8 +105,8 @@ let dft = default_unify_flags let res_pf ?(with_evars=false) ?(with_classes=true) ?(flags=dft ()) clenv = Proofview.Goal.enter { enter = begin fun gl -> - let clenv gl = clenv_unique_resolver ~flags clenv gl in - clenv_refine with_evars ~with_classes (Tacmach.New.of_old clenv (Proofview.Goal.assume gl)) + let clenv = clenv_unique_resolver ~flags clenv gl in + clenv_refine with_evars ~with_classes clenv end } (* [unifyTerms] et [unify] ne semble pas gérer les Meta, en diff --git a/tactics/auto.ml b/tactics/auto.ml index b548f8b92..c8c119aee 100644 --- a/tactics/auto.ml +++ b/tactics/auto.ml @@ -101,9 +101,9 @@ let connect_hint_clenv poly (c, _, ctx) clenv gl = in clenv, c let unify_resolve poly flags ((c : raw_hint), clenv) = - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> let clenv, c = connect_hint_clenv poly c clenv gl in - let clenv = Tacmach.New.of_old (fun gl -> clenv_unique_resolver ~flags clenv gl) gl in + let clenv = clenv_unique_resolver ~flags clenv gl in Clenvtac.clenv_refine false clenv end } @@ -330,7 +330,7 @@ let rec trivial_fail_db dbg mod_delta db_list local_db = end }) in Proofview.Goal.enter { enter = begin fun gl -> - let concl = Tacmach.New.pf_nf_concl gl in + let concl = Tacmach.New.pf_concl gl in let sigma = Tacmach.New.project gl in let secvars = compute_secvars gl in Tacticals.New.tclFIRST @@ -421,7 +421,7 @@ and trivial_resolve sigma dbg mod_delta db_list local_db secvars cl = "nocore" amongst the databases. *) let trivial ?(debug=Off) lems dbnames = - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in let db_list = make_db_list dbnames in @@ -432,7 +432,7 @@ let trivial ?(debug=Off) lems dbnames = end } let full_trivial ?(debug=Off) lems = - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in let db_list = current_pure_db () in @@ -490,7 +490,7 @@ let search d n mod_delta db_list local_db = Tacticals.New.tclORELSE0 (dbg_assumption d) (Tacticals.New.tclORELSE0 (intro_register d (search d n) local_db) ( Proofview.Goal.enter { enter = begin fun gl -> - let concl = Tacmach.New.pf_nf_concl gl in + let concl = Tacmach.New.pf_concl gl in let sigma = Tacmach.New.project gl in let secvars = compute_secvars gl in let d' = incr_dbg d in @@ -506,7 +506,7 @@ let search d n mod_delta db_list local_db = let default_search_depth = ref 5 let delta_auto debug mod_delta n lems dbnames = - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in let db_list = make_db_list dbnames in @@ -529,7 +529,7 @@ let new_auto ?(debug=Off) n = delta_auto debug true n let default_auto = auto !default_search_depth [] [] let delta_full_auto ?(debug=Off) mod_delta n lems = - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in let db_list = current_pure_db () in diff --git a/tactics/autorewrite.ml b/tactics/autorewrite.ml index f43f4b250..e58ec5a31 100644 --- a/tactics/autorewrite.ml +++ b/tactics/autorewrite.ml @@ -92,7 +92,7 @@ type raw_rew_rule = Loc.t * constr Univ.in_universe_context_set * bool * Genarg. let one_base general_rewrite_maybe_in tac_main bas = let lrul = find_rewrites bas in let try_rewrite dir ctx c tc = - Proofview.Goal.nf_s_enter { s_enter = begin fun gl -> + Proofview.Goal.s_enter { s_enter = begin fun gl -> let sigma = Proofview.Goal.sigma gl in let subst, ctx' = Universes.fresh_universe_context_set_instance ctx in let c' = Vars.subst_univs_level_constr subst c in diff --git a/tactics/class_tactics.ml b/tactics/class_tactics.ml index 669d80814..8ada9e6a7 100644 --- a/tactics/class_tactics.ml +++ b/tactics/class_tactics.ml @@ -231,13 +231,13 @@ let e_give_exact flags poly (c,clenv) gl = let unify_e_resolve poly flags = { enter = begin fun gls (c,_,clenv) -> let clenv', c = connect_hint_clenv poly c clenv gls in - let clenv' = Tacmach.New.of_old (clenv_unique_resolver ~flags clenv') gls in + let clenv' = clenv_unique_resolver ~flags clenv' gls in Clenvtac.clenv_refine true ~with_classes:false clenv' end } let unify_resolve poly flags = { enter = begin fun gls (c,_,clenv) -> let clenv', _ = connect_hint_clenv poly c clenv gls in - let clenv' = Tacmach.New.of_old (clenv_unique_resolver ~flags clenv') gls in + let clenv' = clenv_unique_resolver ~flags clenv' gls in Clenvtac.clenv_refine false ~with_classes:false clenv' end } @@ -285,16 +285,16 @@ let clenv_of_prods poly nprods (c, clenv) gl = if Pervasives.(>=) diff 0 then (* Was Some clenv... *) Some (Some diff, - Tacmach.New.of_old (fun gls -> mk_clenv_from_n gls (Some diff) (c,ty)) gl) + mk_clenv_from_n gl (Some diff) (c,ty)) else None let with_prods nprods poly (c, clenv) f = if get_typeclasses_limit_intros () then - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> match clenv_of_prods poly nprods (c, clenv) gl with | None -> Tacticals.New.tclZEROMSG (str"Not enough premisses") | Some (diff, clenv') -> f.enter gl (c, diff, clenv') end } - else Proofview.Goal.nf_enter + else Proofview.Goal.enter { enter = begin fun gl -> if Int.equal nprods 0 then f.enter gl (c, None, clenv) else Tacticals.New.tclZEROMSG (str"Not enough premisses") end } @@ -345,7 +345,7 @@ let rec e_trivial_fail_db only_classes db_list local_db secvars = let open Tacticals.New in let open Tacmach.New in let trivial_fail = - Proofview.Goal.nf_enter { enter = + Proofview.Goal.enter { enter = begin fun gl -> let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in @@ -356,7 +356,7 @@ let rec e_trivial_fail_db only_classes db_list local_db secvars = end } in let trivial_resolve = - Proofview.Goal.nf_enter { enter = + Proofview.Goal.enter { enter = begin fun gl -> let tacs = e_trivial_resolve db_list local_db secvars only_classes (project gl) (pf_concl gl) in @@ -944,7 +944,7 @@ module Search = struct Hint_db.transparent_state cached_hints == st then cached_hints else - let hints = make_hints {it = Goal.goal g; sigma = project g} + let hints = make_hints {it = Goal.goal (Proofview.Goal.assume g); sigma = project g} st only_classes sign in autogoal_cache := (cwd, only_classes, sign, hints); hints @@ -1024,16 +1024,16 @@ module Search = struct (pr_depth (!idx :: info.search_depth) ++ str": trying " ++ Lazy.force pp ++ (if !foundone != true then - str" on" ++ spc () ++ pr_ev s (Proofview.Goal.goal gl) + str" on" ++ spc () ++ pr_ev s (Proofview.Goal.goal (Proofview.Goal.assume gl)) else mt ()))); - let tac_of gls i j = Goal.nf_enter { enter = fun gl' -> + let tac_of gls i j = Goal.enter { enter = fun gl' -> let sigma' = Goal.sigma gl' in let s' = Sigma.to_evar_map sigma' in let _concl = Goal.concl gl' in if !typeclasses_debug > 0 then Feedback.msg_debug (pr_depth (succ j :: i :: info.search_depth) ++ str" : " ++ - pr_ev s' (Proofview.Goal.goal gl')); + pr_ev s' (Proofview.Goal.goal (Proofview.Goal.assume gl'))); let eq c1 c2 = EConstr.eq_constr s' c1 c2 in let hints' = if b && not (Context.Named.equal eq (Goal.hyps gl') (Goal.hyps gl)) @@ -1042,7 +1042,7 @@ module Search = struct make_autogoal_hints info.search_only_classes ~st gl' else info.search_hints in - let dep' = info.search_dep || Proofview.unifiable s' (Goal.goal gl') gls in + let dep' = info.search_dep || Proofview.unifiable s' (Goal.goal (Proofview.Goal.assume gl')) gls in let info' = { search_depth = succ j :: i :: info.search_depth; last_tac = pp; @@ -1059,7 +1059,7 @@ module Search = struct (if !typeclasses_debug > 0 then Feedback.msg_debug (pr_depth (i :: info.search_depth) ++ str": " ++ Lazy.force pp - ++ str" on" ++ spc () ++ pr_ev s (Proofview.Goal.goal gl) + ++ str" on" ++ spc () ++ pr_ev s (Proofview.Goal.goal (Proofview.Goal.assume gl)) ++ str", " ++ int j ++ str" subgoal(s)" ++ (Option.cata (fun k -> str " in addition to the first " ++ int k) (mt()) k))); @@ -1130,7 +1130,7 @@ module Search = struct else tclONCE (aux (NotApplicableEx,Exninfo.null) poss) let hints_tac hints info kont : unit Proofview.tactic = - Proofview.Goal.nf_enter + Proofview.Goal.enter { enter = fun gl -> hints_tac_gl hints info kont gl } let intro_tac info kont gl = @@ -1150,7 +1150,7 @@ module Search = struct let intro info kont = Proofview.tclBIND Tactics.intro - (fun _ -> Proofview.Goal.nf_enter { enter = fun gl -> intro_tac info kont gl }) + (fun _ -> Proofview.Goal.enter { enter = fun gl -> intro_tac info kont gl }) let rec search_tac hints limit depth = let kont info = @@ -1173,7 +1173,7 @@ module Search = struct unit Proofview.tactic = let open Proofview in let open Proofview.Notations in - let dep = dep || Proofview.unifiable sigma (Goal.goal gl) gls in + let dep = dep || Proofview.unifiable sigma (Goal.goal (Proofview.Goal.assume gl)) gls in let info = make_autogoal ?st only_classes dep (cut_of_hints hints) i gl in search_tac hints depth 1 info @@ -1510,11 +1510,11 @@ let is_ground c gl = if Evarutil.is_ground_term (project gl) c then tclIDTAC gl else tclFAIL 0 (str"Not ground") gl -let autoapply c i gl = +let autoapply c i = Proofview.Goal.enter { enter = begin fun gl -> let flags = auto_unif_flags Evar.Set.empty (Hints.Hint_db.transparent_state (Hints.searchtable_map i)) in - let cty = pf_unsafe_type_of gl c in + let cty = Tacmach.New.pf_unsafe_type_of gl c in let ce = mk_clenv_from gl (c,cty) in - let tac = { enter = fun gl -> (unify_e_resolve false flags).enter gl - ((c,cty,Univ.ContextSet.empty),0,ce) } in - Proofview.V82.of_tactic (Proofview.Goal.nf_enter tac) gl + (unify_e_resolve false flags).enter gl + ((c,cty,Univ.ContextSet.empty),0,ce) +end } diff --git a/tactics/class_tactics.mli b/tactics/class_tactics.mli index 171b5c4ea..8855093ee 100644 --- a/tactics/class_tactics.mli +++ b/tactics/class_tactics.mli @@ -31,7 +31,7 @@ val not_evar : constr -> unit Proofview.tactic val is_ground : constr -> tactic -val autoapply : constr -> Hints.hint_db_name -> tactic +val autoapply : constr -> Hints.hint_db_name -> unit Proofview.tactic module Search : sig val eauto_tac : diff --git a/tactics/contradiction.ml b/tactics/contradiction.ml index 0e28aa980..63f923dfd 100644 --- a/tactics/contradiction.ml +++ b/tactics/contradiction.ml @@ -110,7 +110,7 @@ let is_negation_of env sigma typ t = | _ -> false let contradiction_term (c,lbind as cl) = - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> let sigma = Tacmach.New.project gl in let env = Proofview.Goal.env gl in let type_of = Tacmach.New.pf_unsafe_type_of gl in diff --git a/tactics/eauto.ml b/tactics/eauto.ml index 7453fff5c..14082bb8d 100644 --- a/tactics/eauto.ml +++ b/tactics/eauto.ml @@ -112,13 +112,12 @@ open Auto let priority l = List.map snd (List.filter (fun (pr,_) -> Int.equal pr 0) l) let unify_e_resolve poly flags (c,clenv) = - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> let clenv', c = connect_hint_clenv poly c clenv gl in - Proofview.V82.tactic - (fun gls -> - let clenv' = clenv_unique_resolver ~flags clenv' gls in - tclTHEN (Refiner.tclEVARUNIVCONTEXT (Evd.evar_universe_context clenv'.evd)) - (Proofview.V82.of_tactic (Tactics.Simple.eapply c)) gls) + let clenv' = clenv_unique_resolver ~flags clenv' gl in + Proofview.tclTHEN + (Proofview.Unsafe.tclEVARUNIVCONTEXT (Evd.evar_universe_context clenv'.evd)) + (Tactics.Simple.eapply c) end } let hintmap_of sigma secvars hdc concl = @@ -139,7 +138,7 @@ let e_exact poly flags (c,clenv) = end } let rec e_trivial_fail_db db_list local_db = - let next = Proofview.Goal.nf_enter { enter = begin fun gl -> + let next = Proofview.Goal.enter { enter = begin fun gl -> let d = Tacmach.New.pf_last_hyp gl in let hintl = make_resolve_hyp (Tacmach.New.pf_env gl) (Tacmach.New.project gl) d in e_trivial_fail_db db_list (Hint_db.add_list (Tacmach.New.pf_env gl) (Tacmach.New.project gl) hintl local_db) @@ -149,7 +148,7 @@ let rec e_trivial_fail_db db_list local_db = let tacl = registered_e_assumption :: (Tacticals.New.tclTHEN Tactics.intro next) :: - (List.map fst (e_trivial_resolve (Tacmach.New.project gl) db_list local_db secvars (Tacmach.New.pf_nf_concl gl))) + (List.map fst (e_trivial_resolve (Tacmach.New.project gl) db_list local_db secvars (Tacmach.New.pf_concl gl))) in Tacticals.New.tclFIRST (List.map Tacticals.New.tclCOMPLETE tacl) end } @@ -501,7 +500,7 @@ let unfold_head env sigma (ids, csts) c = in aux c let autounfold_one db cl = - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in let concl = Proofview.Goal.concl gl in diff --git a/tactics/elim.ml b/tactics/elim.ml index a4158f821..e37ec6bce 100644 --- a/tactics/elim.ml +++ b/tactics/elim.ml @@ -133,7 +133,7 @@ let induction_trailer abs_i abs_j bargs = (tclDO (abs_j - abs_i) intro) (onLastHypId (fun id -> - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> let idty = pf_unsafe_type_of gl (mkVar id) in let fvty = global_vars (pf_env gl) (project gl) idty in let possible_bring_hyps = @@ -155,7 +155,7 @@ let induction_trailer abs_i abs_j bargs = )) let double_ind h1 h2 = - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> let abs_i = depth_of_quantified_hypothesis true h1 gl in let abs_j = depth_of_quantified_hypothesis true h2 gl in let abs = diff --git a/tactics/eqdecide.ml b/tactics/eqdecide.ml index df60f2c66..bac3980d2 100644 --- a/tactics/eqdecide.ml +++ b/tactics/eqdecide.ml @@ -176,7 +176,7 @@ let solveEqBranch rectype = Proofview.tclORELSE begin Proofview.Goal.enter { enter = begin fun gl -> - let concl = pf_nf_concl gl in + let concl = pf_concl gl in let sigma = project gl in match_eqdec sigma concl >>= fun (eqonleft,op,lhs,rhs,_) -> let (mib,mip) = Global.lookup_inductive rectype in @@ -202,7 +202,7 @@ let decideGralEquality = Proofview.tclORELSE begin Proofview.Goal.enter { enter = begin fun gl -> - let concl = pf_nf_concl gl in + let concl = pf_concl gl in let sigma = project gl in match_eqdec sigma concl >>= fun (eqonleft,_,c1,c2,typ) -> let headtyp = hd_app sigma (pf_compute gl typ) in diff --git a/tactics/equality.ml b/tactics/equality.ml index d9b668517..6fcf529c2 100644 --- a/tactics/equality.ml +++ b/tactics/equality.ml @@ -306,7 +306,7 @@ let general_elim_clause with_evars frzevars tac cls c t l l2r elim = else instantiate_lemma_all frzevars gl c t l l2r concl in let typ = match cls with - | None -> pf_nf_concl gl + | None -> pf_concl gl | Some id -> pf_get_hyp_typ id (Proofview.Goal.assume gl) in let cs = instantiate_lemma typ in @@ -406,7 +406,7 @@ let type_of_clause cls gl = match cls with | Some id -> pf_get_hyp_typ id gl let leibniz_rewrite_ebindings_clause cls lft2rgt tac c t l with_evars frzevars dep_proof_ok hdcncl = - Proofview.Goal.nf_s_enter { s_enter = begin fun gl -> + Proofview.Goal.s_enter { s_enter = begin fun gl -> let evd = Sigma.to_evar_map (Proofview.Goal.sigma gl) in let isatomic = isProd evd (whd_zeta evd hdcncl) in let dep_fun = if isatomic then dependent else dependent_no_evar in @@ -1009,7 +1009,7 @@ let discr_positions env sigma (lbeq,eqn,(t,t1,t2)) eq_clause cpath dirn = let discrEq (lbeq,_,(t,t1,t2) as u) eq_clause = let sigma = eq_clause.evd in - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> let env = Proofview.Goal.env gl in match find_positions env sigma t1 t2 with | Inr _ -> @@ -1019,7 +1019,7 @@ let discrEq (lbeq,_,(t,t1,t2) as u) eq_clause = end } let onEquality with_evars tac (c,lbindc) = - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> let type_of = pf_unsafe_type_of gl in let reduce_to_quantified_ind = pf_apply Tacred.reduce_to_quantified_ind gl in let t = type_of c in @@ -1034,7 +1034,7 @@ let onEquality with_evars tac (c,lbindc) = end } let onNegatedEquality with_evars tac = - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> let sigma = Tacmach.New.project gl in let ccl = Proofview.Goal.concl gl in let env = Proofview.Goal.env gl in @@ -1302,7 +1302,7 @@ let eq_dec_scheme_kind_name = ref (fun _ -> failwith "eq_dec_scheme undefined") let set_eq_dec_scheme_kind k = eq_dec_scheme_kind_name := (fun _ -> k) let inject_if_homogenous_dependent_pair ty = - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> try let sigma = Tacmach.New.project gl in let eq,u,(t,t1,t2) = find_this_eq_data_decompose gl ty in @@ -1458,7 +1458,7 @@ let injConcl = injClause None false None let injHyp clear_flag id = injClause None false (Some (clear_flag,ElimOnIdent (Loc.ghost,id))) let decompEqThen ntac (lbeq,_,(t,t1,t2) as u) clause = - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> let sigma = clause.evd in let env = Proofview.Goal.env gl in match find_positions env sigma t1 t2 with @@ -1567,7 +1567,7 @@ let subst_tuple_term env sigma dep_pair1 dep_pair2 b = (* on for further iterated sigma-tuples *) let cutSubstInConcl l2r eqn = - Proofview.Goal.nf_s_enter { s_enter = begin fun gl -> + Proofview.Goal.s_enter { s_enter = begin fun gl -> let env = Proofview.Goal.env gl in let sigma = Proofview.Goal.sigma gl in let (lbeq,u,(t,e1,e2)) = find_eq_data_decompose gl eqn in @@ -1586,7 +1586,7 @@ let cutSubstInConcl l2r eqn = end } let cutSubstInHyp l2r eqn id = - Proofview.Goal.nf_s_enter { s_enter = begin fun gl -> + Proofview.Goal.s_enter { s_enter = begin fun gl -> let env = Proofview.Goal.env gl in let sigma = Proofview.Goal.sigma gl in let (lbeq,u,(t,e1,e2)) = find_eq_data_decompose gl eqn in @@ -1812,7 +1812,7 @@ let subst_all ?(flags=default_subst_tactic_flags ()) () = Proofview.tclUNIT () end } in - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> let ids = find_equations gl in tclMAP process ids end } @@ -1822,7 +1822,7 @@ let subst_all ?(flags=default_subst_tactic_flags ()) () = (* Old implementation, not able to manage configurations like a=b, a=t, or situations like "a = S b, b = S a", or also accidentally unfolding let-ins *) - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> let sigma = project gl in let find_eq_data_decompose = find_eq_data_decompose gl in let test (_,c) = @@ -1877,7 +1877,7 @@ let rewrite_assumption_cond cond_eq_term cl = with | Failure _ | UserError _ -> arec rest gl end in - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> let gl = Proofview.Goal.lift gl Sigma.Unsafe.le in let hyps = Proofview.Goal.hyps gl in arec hyps gl diff --git a/tactics/inv.ml b/tactics/inv.ml index 632a29721..904a17417 100644 --- a/tactics/inv.ml +++ b/tactics/inv.ml @@ -273,7 +273,7 @@ Nota: with Inversion_clear, only four useless hypotheses let generalizeRewriteIntros as_mode tac depids id = Proofview.tclENV >>= fun env -> - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> let dids = dependent_hyps env id depids gl in let reintros = if as_mode then intros_replacing else intros_possibly_replacing in (tclTHENLIST @@ -342,7 +342,7 @@ let projectAndApply as_mode thin avoid id eqname names depids = (if thin then clear [id] else (remember_first_eq id eqname; tclIDTAC)) in let substHypIfVariable tac id = - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> let sigma = project gl in (** We only look at the type of hypothesis "id" *) let hyp = pf_nf_evar gl (pf_get_hyp_typ id (Proofview.Goal.assume gl)) in @@ -378,7 +378,7 @@ let projectAndApply as_mode thin avoid id eqname names depids = id let nLastDecls i tac = - Proofview.Goal.nf_enter { enter = begin fun gl -> tac (nLastDecls gl i) end } + Proofview.Goal.enter { enter = begin fun gl -> tac (nLastDecls gl i) end } (* Introduction of the equations on arguments othin: discriminates Simple Inversion, Inversion and Inversion_clear @@ -386,7 +386,7 @@ let nLastDecls i tac = Some thin: the equations are rewritten, and cleared if thin is true *) let rewrite_equations as_mode othin neqns names ba = - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> let (depids,nodepids) = split_dep_and_nodep ba.Tacticals.assums gl in let first_eq = ref MoveLast in let avoid = if as_mode then List.map NamedDecl.get_id nodepids else [] in @@ -436,7 +436,7 @@ let rewrite_equations_tac as_mode othin id neqns names ba = tac let raw_inversion inv_kind id status names = - Proofview.Goal.nf_s_enter { s_enter = begin fun gl -> + Proofview.Goal.s_enter { s_enter = begin fun gl -> let sigma = Proofview.Goal.sigma gl in let sigma = Sigma.to_evar_map sigma in let env = Proofview.Goal.env gl in @@ -517,14 +517,14 @@ let dinv_clear_tac id = dinv FullInversionClear None None (NamedHyp id) * back to their places in the hyp-list. *) let invIn k names ids id = - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> let hyps = List.map (fun id -> pf_get_hyp id gl) ids in let concl = Proofview.Goal.concl gl in let sigma = project gl in let nb_prod_init = nb_prod sigma concl in let intros_replace_ids = Proofview.Goal.enter { enter = begin fun gl -> - let concl = pf_nf_concl gl in + let concl = pf_concl gl in let sigma = project gl in let nb_of_new_hyp = nb_prod sigma concl - (List.length hyps + nb_prod_init) diff --git a/tactics/leminv.ml b/tactics/leminv.ml index d864e547c..daa962f1d 100644 --- a/tactics/leminv.ml +++ b/tactics/leminv.ml @@ -262,7 +262,8 @@ let add_inversion_lemma_exn na com comsort bool tac = let lemInv id c gls = try - let clause = mk_clenv_type_of gls c in + let open Tacmach in + let clause = mk_clenv_from_env (pf_env gls) (project gls) None (c, pf_unsafe_type_of gls c) in let clause = clenv_constrain_last_binding (EConstr.mkVar id) clause in Proofview.V82.of_tactic (Clenvtac.res_pf clause ~flags:(Unification.elim_flags ()) ~with_evars:false) gls with @@ -277,7 +278,7 @@ let lemInv id c gls = let lemInv_gen id c = try_intros_until (fun id -> Proofview.V82.tactic (lemInv id c)) id let lemInvIn id c ids = - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> let hyps = List.map (fun id -> pf_get_hyp id gl) ids in let intros_replace_ids = let concl = Proofview.Goal.concl gl in diff --git a/tactics/tacticals.ml b/tactics/tacticals.ml index 94f22f903..27c1987a0 100644 --- a/tactics/tacticals.ml +++ b/tactics/tacticals.ml @@ -267,40 +267,6 @@ let pf_with_evars glsev k gls = let pf_constr_of_global gr k = pf_with_evars (fun gls -> on_snd EConstr.of_constr (pf_apply Evd.fresh_global gls gr)) k -(* computing the case/elim combinators *) - -let gl_make_elim ind gl = - let gr = Indrec.lookup_eliminator (fst ind) (elimination_sort_of_goal gl) in - let (sigma, c) = pf_apply Evd.fresh_global gl gr in - (sigma, EConstr.of_constr c) - -let gl_make_case_dep ind gl = - let sigma = Sigma.Unsafe.of_evar_map (Tacmach.project gl) in - let Sigma (r, sigma, _) = Indrec.build_case_analysis_scheme (pf_env gl) sigma ind true - (elimination_sort_of_goal gl) - in - (Sigma.to_evar_map sigma, EConstr.of_constr r) - -let gl_make_case_nodep ind gl = - let sigma = Sigma.Unsafe.of_evar_map (Tacmach.project gl) in - let Sigma (r, sigma, _) = Indrec.build_case_analysis_scheme (pf_env gl) sigma ind false - (elimination_sort_of_goal gl) - in - (Sigma.to_evar_map sigma, EConstr.of_constr r) - -let make_elim_branch_assumptions ba gl = - let assums = - try List.rev (List.firstn ba.nassums (pf_hyps gl)) - with Failure _ -> anomaly (Pp.str "make_elim_branch_assumptions") in - { ba = ba; assums = assums } - -let elim_on_ba tac ba gl = tac (make_elim_branch_assumptions ba gl) gl - -let make_case_branch_assumptions = make_elim_branch_assumptions - -let case_on_ba tac ba gl = tac (make_case_branch_assumptions ba gl) gl - - (** Tacticals of Ltac defined directly in term of Proofview *) module New = struct open Proofview @@ -534,7 +500,7 @@ module New = struct Proofview.Unsafe.tclEVARS sigma <*> tac >>= check_evars_if let tclDELAYEDWITHHOLES check x tac = - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> let env = Proofview.Goal.env gl in let sigma = Proofview.Goal.sigma gl in let Sigma (x, sigma, _) = x.delayed env sigma in @@ -578,13 +544,13 @@ module New = struct let onLastHyp = onNthHyp 1 let onNthDecl m tac = - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> Proofview.tclUNIT (nthDecl m gl) >>= tac end } let onLastDecl = onNthDecl 1 let ifOnHyp pred tac1 tac2 id = - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> let typ = Tacmach.New.pf_get_hyp_typ id gl in if pred (id,typ) then tac1 id @@ -592,7 +558,7 @@ module New = struct tac2 id end } - let onHyps find tac = Proofview.Goal.nf_enter { enter = begin fun gl -> tac (find.enter gl) end } + let onHyps find tac = Proofview.Goal.enter { enter = begin fun gl -> tac (find.enter gl) end } let afterHyp id tac = Proofview.Goal.enter { enter = begin fun gl -> @@ -625,13 +591,13 @@ module New = struct (* c should be of type A1->.. An->B with B an inductive definition *) let general_elim_then_using mk_elim isrec allnames tac predicate ind (c, t) = - Proofview.Goal.nf_enter { enter = begin fun gl -> - let sigma, elim = Tacmach.New.of_old (mk_elim ind) gl in + Proofview.Goal.enter { enter = begin fun gl -> + let sigma, elim = (mk_elim ind).enter gl in Proofview.tclTHEN (Proofview.Unsafe.tclEVARS sigma) - (Proofview.Goal.nf_enter { enter = begin fun gl -> - let indclause = Tacmach.New.of_old (fun gl -> mk_clenv_from gl (c, t)) gl in + (Proofview.Goal.enter { enter = begin fun gl -> + let indclause = mk_clenv_from gl (c, t) in (* applying elimination_scheme just a little modified *) - let elimclause = Tacmach.New.of_old (fun gls -> mk_clenv_from gls (elim,Tacmach.New.pf_unsafe_type_of gl elim)) gl in + let elimclause = mk_clenv_from gl (elim,Tacmach.New.pf_unsafe_type_of gl elim) in let indmv = match EConstr.kind elimclause.evd (last_arg elimclause.evd elimclause.templval.Evd.rebus) with | Meta mv -> mv @@ -660,7 +626,7 @@ module New = struct | None -> elimclause' | Some p -> clenv_unify ~flags Reduction.CONV (mkMeta pmv) p elimclause' in - let clenv' = Tacmach.New.of_old (clenv_unique_resolver ~flags elimclause') gl in + let clenv' = clenv_unique_resolver ~flags elimclause' gl in let after_tac i = let (hd,largs) = decompose_app clenv'.evd clenv'.templtyp.Evd.rebus in let ba = { branchsign = branchsigns.(i); @@ -679,8 +645,64 @@ module New = struct (Proofview.tclEXTEND [] tclIDTAC branchtacs) end }) end } + let elimination_sort_of_goal gl = + (** Retyping will expand evars anyway. *) + let c = Proofview.Goal.concl (Goal.assume gl) in + pf_apply Retyping.get_sort_family_of gl c + + let elimination_sort_of_hyp id gl = + (** Retyping will expand evars anyway. *) + let c = pf_get_hyp_typ id (Goal.assume gl) in + pf_apply Retyping.get_sort_family_of gl c + + let elimination_sort_of_clause id gl = match id with + | None -> elimination_sort_of_goal gl + | Some id -> elimination_sort_of_hyp id gl + + (* computing the case/elim combinators *) + + let gl_make_elim ind = { enter = begin fun gl -> + let gr = Indrec.lookup_eliminator (fst ind) (elimination_sort_of_goal gl) in + let (sigma, c) = pf_apply Evd.fresh_global gl gr in + (sigma, EConstr.of_constr c) + end } + + let gl_make_case_dep ind = { enter = begin fun gl -> + let sigma = Sigma.Unsafe.of_evar_map (project gl) in + let Sigma (r, sigma, _) = Indrec.build_case_analysis_scheme (pf_env gl) sigma ind true + (elimination_sort_of_goal gl) + in + (Sigma.to_evar_map sigma, EConstr.of_constr r) + end } + + let gl_make_case_nodep ind = { enter = begin fun gl -> + let sigma = Sigma.Unsafe.of_evar_map (project gl) in + let Sigma (r, sigma, _) = Indrec.build_case_analysis_scheme (pf_env gl) sigma ind false + (elimination_sort_of_goal gl) + in + (Sigma.to_evar_map sigma, EConstr.of_constr r) + end } + + let make_elim_branch_assumptions ba hyps = + let assums = + try List.rev (List.firstn ba.nassums hyps) + with Failure _ -> anomaly (Pp.str "make_elim_branch_assumptions") in + { ba = ba; assums = assums } + + let elim_on_ba tac ba = + Proofview.Goal.enter { enter = begin fun gl -> + let branches = make_elim_branch_assumptions ba (Proofview.Goal.hyps gl) in + tac branches + end } + + let case_on_ba tac ba = + Proofview.Goal.enter { enter = begin fun gl -> + let branches = make_elim_branch_assumptions ba (Proofview.Goal.hyps gl) in + tac branches + end } + let elimination_then tac c = - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> let (ind,t) = pf_reduce_to_quantified_ind gl (pf_unsafe_type_of gl c) in let isrec,mkelim = match (Global.lookup_mind (fst (fst ind))).mind_record with @@ -696,34 +718,8 @@ module New = struct let case_nodep_then_using = general_elim_then_using gl_make_case_nodep false - let elim_on_ba tac ba = - Proofview.Goal.nf_enter { enter = begin fun gl -> - let branches = Tacmach.New.of_old (make_elim_branch_assumptions ba) gl in - tac branches - end } - - let case_on_ba tac ba = - Proofview.Goal.nf_enter { enter = begin fun gl -> - let branches = Tacmach.New.of_old (make_case_branch_assumptions ba) gl in - tac branches - end } - - let elimination_sort_of_goal gl = - (** Retyping will expand evars anyway. *) - let c = Proofview.Goal.concl (Goal.assume gl) in - pf_apply Retyping.get_sort_family_of gl c - - let elimination_sort_of_hyp id gl = - (** Retyping will expand evars anyway. *) - let c = pf_get_hyp_typ id (Goal.assume gl) in - pf_apply Retyping.get_sort_family_of gl c - - let elimination_sort_of_clause id gl = match id with - | None -> elimination_sort_of_goal gl - | Some id -> elimination_sort_of_hyp id gl - let pf_constr_of_global ref tac = - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in let (sigma, c) = Evd.fresh_global env sigma ref in diff --git a/tactics/tacticals.mli b/tactics/tacticals.mli index 4bb745875..c9ff77716 100644 --- a/tactics/tacticals.mli +++ b/tactics/tacticals.mli @@ -137,9 +137,6 @@ val elimination_sort_of_clause : Id.t option -> goal sigma -> sorts_family val pf_with_evars : (goal sigma -> Evd.evar_map * 'a) -> ('a -> tactic) -> tactic val pf_constr_of_global : Globnames.global_reference -> (constr -> tactic) -> tactic -val elim_on_ba : (branch_assumptions -> tactic) -> branch_args -> tactic -val case_on_ba : (branch_assumptions -> tactic) -> branch_args -> tactic - (** Tacticals defined directly in term of Proofview *) (** The tacticals in the module [New] are the tactical of Ltac. Their @@ -240,7 +237,7 @@ module New : sig val onLastHyp : (constr -> unit tactic) -> unit tactic val onLastDecl : (named_declaration -> unit tactic) -> unit tactic - val onHyps : ([ `NF ], named_context) Proofview.Goal.enter -> + val onHyps : ([ `LZ ], named_context) Proofview.Goal.enter -> (named_context -> unit tactic) -> unit tactic val afterHyp : Id.t -> (named_context -> unit tactic) -> unit tactic diff --git a/tactics/tactics.ml b/tactics/tactics.ml index 13ffbc52f..5ad43a7d6 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -515,7 +515,7 @@ let rec check_mutind env sigma k cl = match EConstr.kind sigma (strip_outer_cast | _ -> error "Not enough products." (* Refine as a fixpoint *) -let mutual_fix f n rest j = Proofview.Goal.nf_enter { enter = begin fun gl -> +let mutual_fix f n rest j = Proofview.Goal.enter { enter = begin fun gl -> let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in let concl = Proofview.Goal.concl gl in @@ -571,7 +571,7 @@ let rec check_is_mutcoind env sigma cl = error "All methods must construct elements in coinductive types." (* Refine as a cofixpoint *) -let mutual_cofix f others j = Proofview.Goal.nf_enter { enter = begin fun gl -> +let mutual_cofix f others j = Proofview.Goal.enter { enter = begin fun gl -> let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in let concl = Proofview.Goal.concl gl in @@ -697,12 +697,12 @@ let bind_red_expr_occurrences occs nbcl redexp = certain hypothesis *) let reduct_in_concl (redfun,sty) = - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> convert_concl_no_check (Tacmach.New.pf_apply redfun gl (Tacmach.New.pf_concl gl)) sty end } let reduct_in_hyp ?(check=false) redfun (id,where) = - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> convert_hyp ~check (pf_reduce_decl redfun where (Tacmach.New.pf_get_hyp id gl) gl) end } @@ -731,14 +731,14 @@ let pf_e_reduce_decl redfun where decl gl = Sigma (LocalDef (id, b', ty'), sigma, p +> q) let e_reduct_in_concl ~check (redfun, sty) = - Proofview.Goal.nf_s_enter { s_enter = begin fun gl -> + Proofview.Goal.s_enter { s_enter = begin fun gl -> let sigma = Proofview.Goal.sigma gl in let Sigma (c', sigma, p) = redfun.e_redfun (Tacmach.New.pf_env gl) sigma (Tacmach.New.pf_concl gl) in Sigma (convert_concl ~check c' sty, sigma, p) end } let e_reduct_in_hyp ?(check=false) redfun (id, where) = - Proofview.Goal.nf_s_enter { s_enter = begin fun gl -> + Proofview.Goal.s_enter { s_enter = begin fun gl -> let Sigma (decl', sigma, p) = pf_e_reduce_decl redfun where (Tacmach.New.pf_get_hyp id gl) gl in Sigma (convert_hyp ~check decl', sigma, p) end } @@ -1112,7 +1112,7 @@ let depth_of_quantified_hypothesis red h gl = str".") let intros_until_gen red h = - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> let n = depth_of_quantified_hypothesis red h gl in Tacticals.New.tclDO n (if red then introf else intro) end } @@ -1226,7 +1226,7 @@ let cut c = Proofview.Goal.enter { enter = begin fun gl -> let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in - let concl = Tacmach.New.pf_nf_concl gl in + let concl = Proofview.Goal.concl gl in let is_sort = try (** Backward compat: ensure that [c] is well-typed. *) @@ -1360,7 +1360,7 @@ let enforce_prop_bound_names rename tac = mkLetIn (na,c,t,aux (push_rel (LocalDef (na,c,t)) env) sigma (i-1) t') | _ -> assert false in let rename_branch i = - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in let t = Proofview.Goal.concl gl in @@ -1438,7 +1438,7 @@ let general_elim with_evars clear_flag (c, lbindc) elim = (* Case analysis tactics *) let general_case_analysis_in_context with_evars clear_flag (c,lbindc) = - Proofview.Goal.nf_s_enter { s_enter = begin fun gl -> + Proofview.Goal.s_enter { s_enter = begin fun gl -> let sigma = Proofview.Goal.sigma gl in let env = Proofview.Goal.env gl in let concl = Proofview.Goal.concl gl in @@ -1629,7 +1629,7 @@ let make_projection env sigma params cstr sign elim i n c u = in elim let descend_in_conjunctions avoid tac (err, info) c = - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in try @@ -1676,7 +1676,7 @@ let descend_in_conjunctions avoid tac (err, info) c = (****************************************************) let solve_remaining_apply_goals = - Proofview.Goal.nf_s_enter { s_enter = begin fun gl -> + Proofview.Goal.s_enter { s_enter = begin fun gl -> let sigma = Proofview.Goal.sigma gl in if !apply_solve_class_goals then try @@ -1701,7 +1701,7 @@ let tclORELSEOPT t k = | Some tac -> tac) let general_apply with_delta with_destruct with_evars clear_flag (loc,(c,lbind : EConstr.constr with_bindings)) = - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> let concl = Proofview.Goal.concl gl in let sigma = Tacmach.New.project gl in let flags = @@ -1854,7 +1854,7 @@ let apply_in_once_main flags innerclause env sigma (d,lbind) = let apply_in_once sidecond_first with_delta with_destruct with_evars naming id (clear_flag,(loc,(d,lbind))) tac = let open Context.Rel.Declaration in - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in let flags = @@ -1915,7 +1915,7 @@ let apply_in_delayed_once sidecond_first with_delta with_destruct with_evars nam *) let cut_and_apply c = - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> let sigma = Tacmach.New.project gl in match EConstr.kind sigma (Tacmach.New.pf_hnf_constr gl (Tacmach.New.pf_unsafe_type_of gl c)) with | Prod (_,c1,c2) when Vars.noccurn sigma 1 c2 -> @@ -1969,7 +1969,7 @@ let native_cast_no_check c = cast_no_check Term.NATIVEcast c let exact_proof c = let open Tacmach.New in - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> Refine.refine { run = begin fun sigma -> let sigma = Sigma.to_evar_map sigma in let (c, ctx) = Constrintern.interp_casted_constr (pf_env gl) sigma c (pf_concl gl) in @@ -2005,7 +2005,7 @@ let assumption = let hyps = Proofview.Goal.hyps gl in arec gl true hyps end } in - Proofview.Goal.nf_enter assumption_tac + Proofview.Goal.enter assumption_tac (*****************************************************************) (* Modification of a local context *) @@ -2110,7 +2110,7 @@ let rec intros_clearing = function (* Keeping only a few hypotheses *) let keep hyps = - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> Proofview.tclENV >>= fun env -> let ccl = Proofview.Goal.concl gl in let sigma = Tacmach.New.project gl in @@ -2158,7 +2158,7 @@ let bring_hyps hyps = Proofview.Goal.enter { enter = begin fun gl -> let env = Proofview.Goal.env gl in let store = Proofview.Goal.extra gl in - let concl = Tacmach.New.pf_nf_concl gl in + let concl = Tacmach.New.pf_concl gl in let newcl = List.fold_right mkNamedProd_or_LetIn hyps concl in let args = Array.of_list (Context.Named.to_instance mkVar hyps) in Refine.refine { run = begin fun sigma -> @@ -2192,7 +2192,7 @@ let check_number_of_constructors expctdnumopt i nconstr = let constructor_tac with_evars expctdnumopt i lbind = Proofview.Goal.s_enter { s_enter = begin fun gl -> let sigma = Proofview.Goal.sigma gl in - let cl = Tacmach.New.pf_nf_concl gl in + let cl = Tacmach.New.pf_concl gl in let reduce_to_quantified_ind = Tacmach.New.pf_apply Tacred.reduce_to_quantified_ind gl in @@ -2231,7 +2231,7 @@ let any_constructor with_evars tacopt = let t = match tacopt with None -> Proofview.tclUNIT () | Some t -> t in let tac i = Tacticals.New.tclTHEN (constructor_tac with_evars None i NoBindings) t in Proofview.Goal.enter { enter = begin fun gl -> - let cl = Tacmach.New.pf_nf_concl gl in + let cl = Tacmach.New.pf_concl gl in let reduce_to_quantified_ind = Tacmach.New.pf_apply Tacred.reduce_to_quantified_ind gl in @@ -2291,7 +2291,7 @@ let my_find_eq_data_decompose gl t = | Constr_matching.PatternMatchingFailure -> None let intro_decomp_eq loc l thin tac id = - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> let c = mkVar id in let t = Tacmach.New.pf_unsafe_type_of gl c in let _,t = Tacmach.New.pf_reduce_to_quantified_ind gl t in @@ -2702,7 +2702,7 @@ let mkletin_goal env sigma store with_eq dep (id,lastlhyp,ccl,c) ty = Sigma (mkNamedLetIn id c t x, sigma, p) let letin_tac with_eq id c ty occs = - Proofview.Goal.nf_s_enter { s_enter = begin fun gl -> + Proofview.Goal.s_enter { s_enter = begin fun gl -> let sigma = Proofview.Goal.sigma gl in let env = Proofview.Goal.env gl in let ccl = Proofview.Goal.concl gl in @@ -2719,7 +2719,7 @@ let letin_tac with_eq id c ty occs = end } let letin_pat_tac with_eq id c occs = - Proofview.Goal.nf_s_enter { s_enter = begin fun gl -> + Proofview.Goal.s_enter { s_enter = begin fun gl -> let sigma = Proofview.Goal.sigma gl in let env = Proofview.Goal.env gl in let ccl = Proofview.Goal.concl gl in @@ -2805,6 +2805,12 @@ let generalize_goal gl i ((occs,c,b),na as o) (cl,sigma) = let sigma, t = Typing.type_of env sigma c in generalize_goal_gen env sigma ids i o t cl +let new_generalize_goal gl i ((occs,c,b),na as o) (cl,sigma) = + let env = Tacmach.New.pf_env gl in + let ids = Tacmach.New.pf_ids_of_hyps gl in + let sigma, t = Typing.type_of env sigma c in + generalize_goal_gen env sigma ids i o t cl + let old_generalize_dep ?(with_let=false) c gl = let env = pf_env gl in let sign = pf_hyps gl in @@ -2849,10 +2855,10 @@ let generalize_dep ?(with_let = false) c = Proofview.V82.tactic (old_generalize_dep ~with_let c) (** *) -let generalize_gen_let lconstr = Proofview.Goal.nf_s_enter { s_enter = begin fun gl -> +let generalize_gen_let lconstr = Proofview.Goal.s_enter { s_enter = begin fun gl -> let env = Proofview.Goal.env gl in let newcl, evd = - List.fold_right_i (Tacmach.New.of_old generalize_goal gl) 0 lconstr + List.fold_right_i (new_generalize_goal gl) 0 lconstr (Tacmach.New.pf_concl gl,Tacmach.New.project gl) in let (evd, _) = Typing.type_of env evd newcl in @@ -3618,14 +3624,15 @@ let is_defined_variable env id = env |> lookup_named id |> is_local_def let abstract_args gl generalize_vars dep id defined f args = + let open Tacmach.New in let open Context.Rel.Declaration in - let sigma = ref (Tacmach.project gl) in - let env = Tacmach.pf_env gl in - let concl = Tacmach.pf_concl gl in + let sigma = ref (Tacmach.New.project gl) in + let env = Tacmach.New.pf_env gl in + let concl = Tacmach.New.pf_concl gl in let dep = dep || local_occur_var !sigma id concl in let avoid = ref [] in let get_id name = - let id = fresh_id !avoid (match name with Name n -> n | Anonymous -> Id.of_string "gen_x") gl in + let id = new_fresh_id !avoid (match name with Name n -> n | Anonymous -> Id.of_string "gen_x") gl in avoid := id :: !avoid; id in (* Build application generalized w.r.t. the argument plus the necessary eqs. @@ -3640,7 +3647,7 @@ let abstract_args gl generalize_vars dep id defined f args = let decl = List.hd rel in RelDecl.get_name decl, RelDecl.get_type decl, c in - let argty = Tacmach.pf_unsafe_type_of gl arg in + let argty = Tacmach.New.pf_unsafe_type_of gl arg in let sigma', ty = Evarsolve.refresh_universes (Some true) env !sigma ty in let () = sigma := sigma' in let lenctx = List.length ctx in @@ -3681,7 +3688,7 @@ let abstract_args gl generalize_vars dep id defined f args = true, mkApp (f', before), after in if dogen then - let tyf' = Tacmach.pf_unsafe_type_of gl f' in + let tyf' = Tacmach.New.pf_unsafe_type_of gl f' in let arity, ctx, ctxenv, c', args, eqs, refls, nogen, vars, env = Array.fold_left aux (tyf',[],env,f',[],[],[],Id.Set.empty,Id.Set.empty,env) args' in @@ -3689,14 +3696,14 @@ let abstract_args gl generalize_vars dep id defined f args = let vars = if generalize_vars then let nogen = Id.Set.add id nogen in - hyps_of_vars (pf_env gl) (project gl) (pf_hyps gl) nogen vars + hyps_of_vars (pf_env gl) (project gl) (Proofview.Goal.hyps gl) nogen vars else [] in let body, c' = if defined then Some c', Retyping.get_type_of ctxenv !sigma c' else None, c' in - let typ = Tacmach.pf_get_hyp_typ gl id in + let typ = Tacmach.New.pf_get_hyp_typ id gl in let tac = make_abstract_generalize (pf_env gl) id typ concl dep ctx body c' eqs args refls in let tac = Proofview.Unsafe.tclEVARS !sigma <*> tac in Some (tac, dep, succ (List.length ctx), vars) @@ -3704,7 +3711,7 @@ let abstract_args gl generalize_vars dep id defined f args = let abstract_generalize ?(generalize_vars=true) ?(force_dep=false) id = let open Context.Named.Declaration in - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> Coqlib.check_required_library Coqlib.jmeq_module_name; let sigma = Tacmach.New.project gl in let (f, args, def, id, oldid) = @@ -3719,7 +3726,7 @@ let abstract_generalize ?(generalize_vars=true) ?(force_dep=false) id = if List.is_empty args then Proofview.tclUNIT () else let args = Array.of_list args in - let newc = Tacmach.New.of_old (fun gl -> abstract_args gl generalize_vars force_dep id def f args) gl in + let newc = abstract_args gl generalize_vars force_dep id def f args in match newc with | None -> Proofview.tclUNIT () | Some (tac, dep, n, vars) -> @@ -3799,7 +3806,7 @@ let specialize_eqs id gl = else tclFAIL 0 (str "Nothing to do in hypothesis " ++ pr_id id) gl -let specialize_eqs id = Proofview.Goal.nf_enter { enter = begin fun gl -> +let specialize_eqs id = Proofview.Goal.enter { enter = begin fun gl -> let msg = str "Specialization not allowed on dependent hypotheses" in Proofview.tclOR (clear [id]) (fun _ -> Tacticals.New.tclZEROMSG msg) >>= fun () -> @@ -4123,7 +4130,7 @@ let recolle_clenv i params args elimclause gl = (* from_n (Some 0) means that x should be taken "as is" without trying to unify (which would lead to trying to apply it to evars if y is a product). *) - let indclause = Tacmach.New.of_old (fun gl -> mk_clenv_from_n gl (Some 0) (x,y)) gl in + let indclause = mk_clenv_from_n gl (Some 0) (x,y) in let elimclause' = clenv_fchain ~with_univs:false i acc indclause in elimclause') (List.rev clauses) @@ -4134,18 +4141,18 @@ let recolle_clenv i params args elimclause gl = produce new ones). Then refine with the resulting term with holes. *) let induction_tac with_evars params indvars elim = - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> let sigma = Tacmach.New.project gl in let ({elimindex=i;elimbody=(elimc,lbindelimc);elimrename=rename},elimt) = elim in let i = match i with None -> index_of_ind_arg sigma elimt | Some i -> i in (* elimclause contains this: (elimc ?i ?j ?k...?l) *) let elimc = contract_letin_in_lam_header sigma elimc in let elimc = mkCast (elimc, DEFAULTcast, elimt) in - let elimclause = pf_apply make_clenv_binding gl (elimc,elimt) lbindelimc in + let elimclause = Tacmach.New.pf_apply make_clenv_binding gl (elimc,elimt) lbindelimc in (* elimclause' is built from elimclause by instanciating all args and params. *) let elimclause' = recolle_clenv i params indvars elimclause gl in (* one last resolution (useless?) *) - let resolved = Tacmach.New.of_old (clenv_unique_resolver ~flags:(elim_flags ()) elimclause') gl in + let resolved = clenv_unique_resolver ~flags:(elim_flags ()) elimclause' gl in enforce_prop_bound_names rename (Clenvtac.clenv_refine with_evars resolved) end } @@ -4158,7 +4165,7 @@ let apply_induction_in_context with_evars hyp0 inhyps elim indvars names induct_ let sigma = Proofview.Goal.sigma gl in let env = Proofview.Goal.env gl in let sigma = Sigma.to_evar_map sigma in - let concl = Tacmach.New.pf_nf_concl gl in + let concl = Tacmach.New.pf_concl gl in let statuslists,lhyp0,toclear,deps,avoid,dep_in_hyps = cook_sign hyp0 inhyps indvars env sigma in let dep_in_concl = Option.cata (fun id -> occur_var env sigma id concl) false hyp0 in let dep = dep_in_hyps || dep_in_concl in @@ -4212,7 +4219,7 @@ let msg_not_right_number_induction_arguments scheme = must be given, so we help a bit the unifier by making the "pattern" by hand before calling induction_tac *) let induction_without_atomization isrec with_evars elim names lid = - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> let sigma, (indsign,scheme) = get_elim_signature elim (List.hd lid) gl in let nargs_indarg_farg = scheme.nargs + (if scheme.farg_in_concl then 1 else 0) in @@ -4247,7 +4254,7 @@ let induction_without_atomization isrec with_evars elim names lid = (* assume that no occurrences are selected *) let clear_unselected_context id inhyps cls = - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> if occur_var (Tacmach.New.pf_env gl) (Tacmach.New.project gl) id (Tacmach.New.pf_concl gl) && cls.concl_occs == NoOccurrences then user_err @@ -4493,7 +4500,7 @@ let induction_destruct isrec with_evars (lc,elim) = match lc with | [] -> assert false (* ensured by syntax, but if called inside caml? *) | [c,(eqname,names as allnames),cls] -> - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in match elim with @@ -4594,8 +4601,8 @@ let simple_destruct = function *) let elim_scheme_type elim t = - Proofview.Goal.nf_enter { enter = begin fun gl -> - let clause = Tacmach.New.of_old (fun gl -> mk_clenv_type_of gl elim) gl in + Proofview.Goal.enter { enter = begin fun gl -> + let clause = mk_clenv_type_of gl elim in match EConstr.kind clause.evd (last_arg clause.evd clause.templval.rebus) with | Meta mv -> let clause' = @@ -4634,7 +4641,7 @@ let case_type t = let (forward_setoid_reflexivity, setoid_reflexivity) = Hook.make () let maybe_betadeltaiota_concl allowred gl = - let concl = Tacmach.New.pf_nf_concl gl in + let concl = Tacmach.New.pf_concl gl in let sigma = Tacmach.New.project gl in if not allowred then concl else @@ -4891,7 +4898,7 @@ let abstract_subproof id gk tac = let open Tacticals.New in let open Tacmach.New in let open Proofview.Notations in - Proofview.Goal.nf_s_enter { s_enter = begin fun gl -> + Proofview.Goal.s_enter { s_enter = begin fun gl -> let sigma = Proofview.Goal.sigma gl in let current_sign = Global.named_context_val () and global_sign = Proofview.Goal.hyps gl in @@ -4980,7 +4987,7 @@ let tclABSTRACT name_op tac = abstract_subproof s gk tac let unify ?(state=full_transparent_state) x y = - Proofview.Goal.nf_s_enter { s_enter = begin fun gl -> + Proofview.Goal.s_enter { s_enter = begin fun gl -> let sigma = Proofview.Goal.sigma gl in try let core_flags = diff --git a/tactics/tactics.mli b/tactics/tactics.mli index 0087d607d..67e29cf56 100644 --- a/tactics/tactics.mli +++ b/tactics/tactics.mli @@ -29,7 +29,7 @@ open Locus (** {6 General functions. } *) -val is_quantified_hypothesis : Id.t -> ([`NF],'b) Proofview.Goal.t -> bool +val is_quantified_hypothesis : Id.t -> ('a, 'r) Proofview.Goal.t -> bool (** {6 Primitive tactics. } *) @@ -75,7 +75,7 @@ val intros : unit Proofview.tactic (** [depth_of_quantified_hypothesis b h g] returns the index of [h] in the conclusion of goal [g], up to head-reduction if [b] is [true] *) val depth_of_quantified_hypothesis : - bool -> quantified_hypothesis -> ([`NF],'b) Proofview.Goal.t -> int + bool -> quantified_hypothesis -> ('a, 'r) Proofview.Goal.t -> int val intros_until : quantified_hypothesis -> unit Proofview.tactic diff --git a/toplevel/auto_ind_decl.ml b/toplevel/auto_ind_decl.ml index f343cc73d..e90d8a4fd 100644 --- a/toplevel/auto_ind_decl.ml +++ b/toplevel/auto_ind_decl.ml @@ -364,8 +364,8 @@ let do_replace_lb mode lb_scheme_key aavoid narg p q = ))) ) in - Proofview.Goal.nf_enter { enter = begin fun gl -> - let type_of_pq = Tacmach.New.of_old (fun gl -> pf_unsafe_type_of gl p) gl in + Proofview.Goal.enter { enter = begin fun gl -> + let type_of_pq = Tacmach.New.pf_unsafe_type_of gl p in let sigma = Tacmach.New.project gl in let u,v = destruct_ind sigma type_of_pq in let lb_type_of_p = @@ -574,12 +574,10 @@ let compute_bl_tact mode bl_scheme_key ind lnamesparrec nparrec = ( List.map (fun (_,_,sbl,_ ) -> sbl) list_id ) in let fresh_id s gl = - Tacmach.New.of_old begin fun gsig -> - let fresh = fresh_id (!avoid) s gsig in + let fresh = fresh_id_in_env (!avoid) s (Proofview.Goal.env gl) in avoid := fresh::(!avoid); fresh - end gl in - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> let fresh_first_intros = List.map (fun id -> fresh_id id gl) first_intros in let freshn = fresh_id (Id.of_string "x") gl in let freshm = fresh_id (Id.of_string "y") gl in @@ -602,7 +600,7 @@ repeat ( apply andb_prop in z;let z1:= fresh "Z" in destruct z as [z1 z]). Tacticals.New.tclREPEAT ( Tacticals.New.tclTHENLIST [ Simple.apply_in freshz (EConstr.of_constr (andb_prop())); - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> let fresht = fresh_id (Id.of_string "Z") gl in destruct_on_as (EConstr.mkVar freshz) (IntroOrPattern [[dl,IntroNaming (IntroIdentifier fresht); @@ -613,7 +611,7 @@ repeat ( apply andb_prop in z;let z1:= fresh "Z" in destruct z as [z1 z]). Ci a1 ... an = Ci b1 ... bn replace bi with ai; auto || replace bi with ai by apply typeofbi_prod ; auto *) - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> let concl = Proofview.Goal.concl gl in let sigma = Tacmach.New.project gl in match EConstr.kind sigma concl with @@ -720,12 +718,10 @@ let compute_lb_tact mode lb_scheme_key ind lnamesparrec nparrec = ( List.map (fun (_,_,_,slb) -> slb) list_id ) in let fresh_id s gl = - Tacmach.New.of_old begin fun gsig -> - let fresh = fresh_id (!avoid) s gsig in + let fresh = fresh_id_in_env (!avoid) s (Proofview.Goal.env gl) in avoid := fresh::(!avoid); fresh - end gl in - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> let fresh_first_intros = List.map (fun id -> fresh_id id gl) first_intros in let freshn = fresh_id (Id.of_string "x") gl in let freshm = fresh_id (Id.of_string "y") gl in @@ -748,7 +744,7 @@ let compute_lb_tact mode lb_scheme_key ind lnamesparrec nparrec = Tacticals.New.tclTHENLIST [apply (EConstr.of_constr (andb_true_intro())); simplest_split ;Auto.default_auto ] ); - Proofview.Goal.nf_enter { enter = begin fun gls -> + Proofview.Goal.enter { enter = begin fun gls -> let concl = Proofview.Goal.concl gls in let sigma = Tacmach.New.project gl in (* assume the goal to be eq (eq_type ...) = true *) @@ -869,12 +865,10 @@ let compute_dec_tact ind lnamesparrec nparrec = ( List.map (fun (_,_,_,slb) -> slb) list_id ) in let fresh_id s gl = - Tacmach.New.of_old begin fun gsig -> - let fresh = fresh_id (!avoid) s gsig in + let fresh = fresh_id_in_env (!avoid) s (Proofview.Goal.env gl) in avoid := fresh::(!avoid); fresh - end gl in - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> let fresh_first_intros = List.map (fun id -> fresh_id id gl) first_intros in let freshn = fresh_id (Id.of_string "x") gl in let freshm = fresh_id (Id.of_string "y") gl in @@ -905,7 +899,7 @@ let compute_dec_tact ind lnamesparrec nparrec = )) (Tacticals.New.tclTHEN (destruct_on (EConstr.of_constr eqbnm)) Auto.default_auto); - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> let freshH2 = fresh_id (Id.of_string "H") gl in Tacticals.New.tclTHENS (destruct_on_using (EConstr.mkVar freshH) freshH2) [ (* left *) @@ -917,7 +911,7 @@ let compute_dec_tact ind lnamesparrec nparrec = ; (*right *) - Proofview.Goal.nf_enter { enter = begin fun gl -> + Proofview.Goal.enter { enter = begin fun gl -> let freshH3 = fresh_id (Id.of_string "H") gl in Tacticals.New.tclTHENLIST [ simplest_right ; -- cgit v1.2.3 From 594ac9654164e377e8598894019cc4445509d570 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sat, 17 Dec 2016 18:36:59 +0100 Subject: Removing a subtle nf_enter in Class_tactics. The underlying hint mode implementation was not using the evar-insensitive API so that it resulted in strange bugs. --- tactics/class_tactics.ml | 2 +- tactics/hints.ml | 4 ++-- 2 files changed, 3 insertions(+), 3 deletions(-) diff --git a/tactics/class_tactics.ml b/tactics/class_tactics.ml index 8ada9e6a7..8bbef39ad 100644 --- a/tactics/class_tactics.ml +++ b/tactics/class_tactics.ml @@ -1180,7 +1180,7 @@ module Search = struct let search_tac ?(st=full_transparent_state) only_classes dep hints depth = let open Proofview in let tac sigma gls i = - Goal.nf_enter + Goal.enter { enter = fun gl -> search_tac_gl ~st only_classes dep hints depth (succ i) sigma gls gl } in diff --git a/tactics/hints.ml b/tactics/hints.ml index 17c81064d..5aacafd6f 100644 --- a/tactics/hints.ml +++ b/tactics/hints.ml @@ -524,8 +524,8 @@ struct match m with | ModeInput -> not (occur_existential sigma arg) | ModeNoHeadEvar -> - Evarutil.(try ignore(head_evar sigma arg); false - with NoHeadEvar -> true) + (try ignore(head_evar sigma arg); false + with Evarutil.NoHeadEvar -> true) | ModeOutput -> true let matches_mode sigma args mode = -- cgit v1.2.3 From aaf75678a13d9c26341e762ab8e56b957cf4c771 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Mon, 19 Dec 2016 01:30:45 +0100 Subject: Dedicated datatype for aliases in Evarsolve. --- pretyping/evarsolve.ml | 312 ++++++++++++++++++++++++++++------------------- pretyping/evarsolve.mli | 10 +- pretyping/unification.ml | 4 +- 3 files changed, 197 insertions(+), 129 deletions(-) diff --git a/pretyping/evarsolve.ml b/pretyping/evarsolve.ml index 28e63d04b..398f2665e 100644 --- a/pretyping/evarsolve.ml +++ b/pretyping/evarsolve.ml @@ -246,6 +246,47 @@ let noccur_evar env evd evk c = (* Managing chains of local definitons *) (***************************************) +type alias = +| RelAlias of int +| VarAlias of Id.t + +let of_alias = function +| RelAlias n -> mkRel n +| VarAlias id -> mkVar id + +let to_alias sigma c = match EConstr.kind sigma c with +| Rel n -> Some (RelAlias n) +| Var id -> Some (VarAlias id) +| _ -> None + +let is_alias sigma c alias = match EConstr.kind sigma c, alias with +| Var id, VarAlias id' -> Id.equal id id' +| Rel n, RelAlias n' -> Int.equal n n' +| _ -> false + +let eq_alias a b = match a, b with +| RelAlias n, RelAlias m -> Int.equal m n +| VarAlias id1, VarAlias id2 -> Id.equal id1 id2 +| _ -> false + +type aliasing = EConstr.t option * alias list + +let empty_aliasing = None, [] +let make_aliasing c = Some c, [] +let push_alias (alias, l) a = (alias, a :: l) +let lift_aliasing n (alias, l) = + let map a = match a with + | VarAlias _ -> a + | RelAlias m -> RelAlias (m + n) + in + (Option.map (fun c -> lift n c) alias, List.map map l) + +type aliases = { + rel_aliases : aliasing Int.Map.t; + var_aliases : aliasing Id.Map.t; + (** Only contains [VarAlias] *) +} + (* Expand rels and vars that are bound to other rels or vars so that dependencies in variables are canonically associated to the most ancient variable in its family of aliased variables *) @@ -259,10 +300,10 @@ let compute_var_aliases sign sigma = (match EConstr.kind sigma t with | Var id' -> let aliases_of_id = - try Id.Map.find id' aliases with Not_found -> [] in - Id.Map.add id (aliases_of_id@[t]) aliases + try Id.Map.find id' aliases with Not_found -> empty_aliasing in + Id.Map.add id (push_alias aliases_of_id (VarAlias id')) aliases | _ -> - Id.Map.add id [t] aliases) + Id.Map.add id (make_aliasing t) aliases) | LocalAssum _ -> aliases) sign Id.Map.empty @@ -275,14 +316,14 @@ let compute_rel_aliases var_aliases rels sigma = (match EConstr.kind sigma t with | Var id' -> let aliases_of_n = - try Id.Map.find id' var_aliases with Not_found -> [] in - Int.Map.add n (aliases_of_n@[t]) aliases + try Id.Map.find id' var_aliases with Not_found -> empty_aliasing in + Int.Map.add n (push_alias aliases_of_n (VarAlias id')) aliases | Rel p -> let aliases_of_n = - try Int.Map.find (p+n) aliases with Not_found -> [] in - Int.Map.add n (aliases_of_n@[mkRel (p+n)]) aliases + try Int.Map.find (p+n) aliases with Not_found -> empty_aliasing in + Int.Map.add n (push_alias aliases_of_n (RelAlias (p+n))) aliases | _ -> - Int.Map.add n [lift n (mkCast(t,DEFAULTcast,u))] aliases) + Int.Map.add n (make_aliasing (lift n (mkCast(t,DEFAULTcast,u)))) aliases) | LocalAssum _ -> aliases) ) rels @@ -292,37 +333,43 @@ let make_alias_map env sigma = (* We compute the chain of aliases for each var and rel *) let var_aliases = compute_var_aliases (named_context env) sigma in let rel_aliases = compute_rel_aliases var_aliases (rel_context env) sigma in - (var_aliases,rel_aliases) + { var_aliases; rel_aliases } -let lift_aliases n (var_aliases,rel_aliases as aliases) = +let lift_aliases n aliases = if Int.equal n 0 then aliases else - (var_aliases, - Int.Map.fold (fun p l -> Int.Map.add (p+n) (List.map (lift n) l)) - rel_aliases Int.Map.empty) + let rel_aliases = + Int.Map.fold (fun p l -> Int.Map.add (p+n) (lift_aliasing n l)) + aliases.rel_aliases Int.Map.empty + in + { aliases with rel_aliases } -let get_alias_chain_of sigma aliases x = match EConstr.kind sigma x with - | Rel n -> (try Int.Map.find n (snd aliases) with Not_found -> []) - | Var id -> (try Id.Map.find id (fst aliases) with Not_found -> []) - | _ -> [] +let get_alias_chain_of sigma aliases x = match x with + | RelAlias n -> (try Int.Map.find n aliases.rel_aliases with Not_found -> empty_aliasing) + | VarAlias id -> (try Id.Map.find id aliases.var_aliases with Not_found -> empty_aliasing) -let normalize_alias_opt sigma aliases x = +let normalize_alias_opt_alias sigma aliases x = match get_alias_chain_of sigma aliases x with - | [] -> None - | a::_ when isRel sigma a || isVar sigma a -> Some a - | [_] -> None - | _::a::_ -> Some a + | _, [] -> None + | _, a :: _ -> Some a + +let normalize_alias_opt sigma aliases x = match to_alias sigma x with +| None -> None +| Some a -> normalize_alias_opt_alias sigma aliases a let normalize_alias sigma aliases x = - match normalize_alias_opt sigma aliases x with + match normalize_alias_opt_alias sigma aliases x with | Some a -> a | None -> x let normalize_alias_var sigma var_aliases id = - destVar sigma (normalize_alias sigma (var_aliases,Int.Map.empty) (mkVar id)) + let aliases = { var_aliases; rel_aliases = Int.Map.empty } in + match normalize_alias sigma aliases (VarAlias id) with + | VarAlias id -> id + | RelAlias _ -> assert false (** var only aliases to variables *) -let extend_alias sigma decl (var_aliases,rel_aliases) = +let extend_alias sigma decl { var_aliases; rel_aliases } = let rel_aliases = - Int.Map.fold (fun n l -> Int.Map.add (n+1) (List.map (lift 1) l)) + Int.Map.fold (fun n l -> Int.Map.add (n+1) (lift_aliasing 1 l)) rel_aliases Int.Map.empty in let rel_aliases = match decl with @@ -330,36 +377,36 @@ let extend_alias sigma decl (var_aliases,rel_aliases) = (match EConstr.kind sigma t with | Var id' -> let aliases_of_binder = - try Id.Map.find id' var_aliases with Not_found -> [] in - Int.Map.add 1 (aliases_of_binder@[t]) rel_aliases + try Id.Map.find id' var_aliases with Not_found -> empty_aliasing in + Int.Map.add 1 (push_alias aliases_of_binder (VarAlias id')) rel_aliases | Rel p -> let aliases_of_binder = - try Int.Map.find (p+1) rel_aliases with Not_found -> [] in - Int.Map.add 1 (aliases_of_binder@[mkRel (p+1)]) rel_aliases + try Int.Map.find (p+1) rel_aliases with Not_found -> empty_aliasing in + Int.Map.add 1 (push_alias aliases_of_binder (RelAlias (p+1))) rel_aliases | _ -> - Int.Map.add 1 [lift 1 t] rel_aliases) + Int.Map.add 1 (make_aliasing (lift 1 t)) rel_aliases) | LocalAssum _ -> rel_aliases in - (var_aliases, rel_aliases) + { var_aliases; rel_aliases } let expand_alias_once sigma aliases x = match get_alias_chain_of sigma aliases x with - | [] -> None - | l -> Some (List.last l) + | None, [] -> None + | Some a, [] -> Some a + | _, l -> Some (of_alias (List.last l)) let expansions_of_var sigma aliases x = - match get_alias_chain_of sigma aliases x with - | [] -> [x] - | a::_ as l when isRel sigma a || isVar sigma a -> x :: List.rev l - | _::l -> x :: List.rev l + let (_, l) = get_alias_chain_of sigma aliases x in + x :: List.rev l let expansion_of_var sigma aliases x = match get_alias_chain_of sigma aliases x with - | [] -> x - | a::_ -> a + | None, [] -> (false, of_alias x) + | Some a, _ -> (true, a) + | None, a :: _ -> (true, of_alias a) let rec expand_vars_in_term_using sigma aliases t = match EConstr.kind sigma t with - | Rel _ | Var _ -> - normalize_alias sigma aliases t + | Rel n -> of_alias (normalize_alias sigma aliases (RelAlias n)) + | Var id -> of_alias (normalize_alias sigma aliases (VarAlias id)) | _ -> let self aliases c = expand_vars_in_term_using sigma aliases c in map_constr_with_full_binders sigma (extend_alias sigma) self aliases t @@ -371,24 +418,28 @@ let free_vars_and_rels_up_alias_expansion sigma aliases c = let acc3 = ref Int.Set.empty and acc4 = ref Id.Set.empty in let cache_rel = ref Int.Set.empty and cache_var = ref Id.Set.empty in let is_in_cache depth = function - | Rel n -> Int.Set.mem (n-depth) !cache_rel - | Var s -> Id.Set.mem s !cache_var - | _ -> false in + | RelAlias n -> Int.Set.mem (n-depth) !cache_rel + | VarAlias s -> Id.Set.mem s !cache_var + in let put_in_cache depth = function - | Rel n -> cache_rel := Int.Set.add (n-depth) !cache_rel - | Var s -> cache_var := Id.Set.add s !cache_var - | _ -> () in + | RelAlias n -> cache_rel := Int.Set.add (n-depth) !cache_rel + | VarAlias s -> cache_var := Id.Set.add s !cache_var + in let rec frec (aliases,depth) c = match EConstr.kind sigma c with | Rel _ | Var _ as ck -> + let ck = match ck with + | Rel n -> RelAlias n + | Var id -> VarAlias id + | _ -> assert false + in if is_in_cache depth ck then () else begin put_in_cache depth ck; - let c' = expansion_of_var sigma aliases c in - (if c != c' then (* expansion, hence a let-in *) (** FIXME *) - match EConstr.kind sigma c with - | Var id -> acc4 := Id.Set.add id !acc4 - | Rel n -> if n >= depth+1 then acc3 := Int.Set.add (n-depth) !acc3 - | _ -> ()); + let expanded, c' = expansion_of_var sigma aliases ck in + (if expanded then (* expansion, hence a let-in *) + match ck with + | VarAlias id -> acc4 := Id.Set.add id !acc4 + | RelAlias n -> if n >= depth+1 then acc3 := Int.Set.add (n-depth) !acc3); match EConstr.kind sigma c' with | Var id -> acc2 := Id.Set.add id !acc2 | Rel n -> if n >= depth+1 then acc1 := Int.Set.add (n-depth) !acc1 @@ -407,30 +458,33 @@ let free_vars_and_rels_up_alias_expansion sigma aliases c = (* Managing pattern-unification *) (********************************) -let rec expand_and_check_vars sigma aliases = function +let map_all f l = + let rec map_aux f l = match l with | [] -> [] - | a::l when isRel sigma a || isVar sigma a -> - let a = expansion_of_var sigma aliases a in - if isRel sigma a || isVar sigma a then a :: expand_and_check_vars sigma aliases l - else raise Exit - | _ -> - raise Exit - -module Constrhash = Hashtbl.Make - (struct type t = Constr.constr - let equal = Term.eq_constr - let hash = hash_constr - end) - -let constr_list_distinct sigma l = - let visited = Constrhash.create 23 in - let rec loop = function - | h::t -> - let h = EConstr.to_constr sigma h in - if Constrhash.mem visited h then false - else (Constrhash.add visited h h; loop t) - | [] -> true - in loop l + | x :: l -> + match f x with + | None -> raise Exit + | Some y -> y :: map_aux f l + in + try Some (map_aux f l) with Exit -> None + +let expand_and_check_vars sigma aliases l = + let map a = match get_alias_chain_of sigma aliases a with + | None, [] -> Some a + | None, a :: _ -> Some a + | Some _, _ -> None + in + map_all map l + +let alias_distinct l = + let rec check (rels, vars) = function + | [] -> true + | RelAlias n :: l -> + not (Int.Set.mem n rels) && check (Int.Set.add n rels, vars) l + | VarAlias id :: l -> + not (Id.Set.mem id vars) && check (rels, Id.Set.add id vars) l + in + check (Int.Set.empty, Id.Set.empty) l let get_actual_deps evd aliases l t = if occur_meta_or_existential evd t then @@ -439,11 +493,10 @@ let get_actual_deps evd aliases l t = else (* Probably strong restrictions coming from t being evar-closed *) let (fv_rels,fv_ids,_,_) = free_vars_and_rels_up_alias_expansion evd aliases t in - List.filter (fun c -> - match EConstr.kind evd c with - | Var id -> Id.Set.mem id fv_ids - | Rel n -> Int.Set.mem n fv_rels - | _ -> assert false) l + List.filter (function + | VarAlias id -> Id.Set.mem id fv_ids + | RelAlias n -> Int.Set.mem n fv_rels + ) l open Context.Named.Declaration let remove_instance_local_defs evd evk args = @@ -463,34 +516,41 @@ let remove_instance_local_defs evd evk args = (* Check if an applied evar "?X[args] l" is a Miller's pattern *) let find_unification_pattern_args env evd l t = - if List.for_all (fun x -> isRel evd x || isVar evd x) l (* common failure case *) then - let aliases = make_alias_map env evd in - match (try Some (expand_and_check_vars evd aliases l) with Exit -> None) with - | Some l as x when constr_list_distinct evd (get_actual_deps evd aliases l t) -> x - | _ -> None - else - None + let aliases = make_alias_map env evd in + match expand_and_check_vars evd aliases l with + | Some l as x when alias_distinct (get_actual_deps evd aliases l t) -> x + | _ -> None let is_unification_pattern_meta env evd nb m l t = (* Variables from context and rels > nb are implicitly all there *) (* so we need to be a rel <= nb *) - if List.for_all (fun x -> isRel evd x && destRel evd x <= nb) l then - match find_unification_pattern_args env evd l t with + let map a = match EConstr.kind evd a with + | Rel n -> if n <= nb then Some (RelAlias n) else None + | _ -> None + in + match map_all map l with + | Some l -> + begin match find_unification_pattern_args env evd l t with | Some _ as x when not (dependent evd (mkMeta m) t) -> x | _ -> None - else + end + | None -> None let is_unification_pattern_evar env evd (evk,args) l t = - if List.for_all (fun x -> isRel evd x || isVar evd x) l - && noccur_evar env evd evk t - then + match map_all (fun c -> to_alias evd c) l with + | Some l when noccur_evar env evd evk t -> let args = remove_instance_local_defs evd evk args in + let args = map_all (fun c -> to_alias evd c) args in + begin match args with + | None -> None + | Some args -> let n = List.length args in match find_unification_pattern_args env evd (args @ l) t with | Some l -> Some (List.skipn n l) | _ -> None - else None + end + | _ -> None let is_unification_pattern_pure_evar env evd (evk,args) t = let is_ev = is_unification_pattern_evar env evd (evk,args) [] t in @@ -513,16 +573,16 @@ let is_unification_pattern (env,nb) evd f l t = return \y. ?1{x\y} (non constant function if ?1 depends on x) (BB) *) let solve_pattern_eqn env sigma l c = let c' = List.fold_right (fun a c -> - let c' = subst_term sigma (lift 1 a) (lift 1 c) in - match EConstr.kind sigma a with + let c' = subst_term sigma (lift 1 (of_alias a)) (lift 1 c) in + match a with (* Rem: if [a] links to a let-in, do as if it were an assumption *) - | Rel n -> + | RelAlias n -> let open Context.Rel.Declaration in let d = map_constr (lift n) (lookup_rel n env) in mkLambda_or_LetIn d c' - | Var id -> + | VarAlias id -> let d = lookup_named id env in mkNamedLambda_or_LetIn d c' - | _ -> assert false) + ) l c in (* Warning: we may miss some opportunity to eta-reduce more since c' is not in normal form *) @@ -731,15 +791,15 @@ exception NotUniqueInType of (Id.t * evar_projection) list let rec assoc_up_to_alias sigma aliases y yc = function | [] -> raise Not_found | (c,cc,id)::l -> - if EConstr.eq_constr sigma y c then id + if is_alias sigma c y then id else match l with | _ :: _ -> assoc_up_to_alias sigma aliases y yc l | [] -> (* Last chance, we reason up to alias conversion *) match (normalize_alias_opt sigma aliases c) with - | Some cc when EConstr.eq_constr sigma yc cc -> id - | _ -> if EConstr.eq_constr sigma yc c then id else raise Not_found + | Some cc when eq_alias yc cc -> id + | _ -> if is_alias sigma c yc then id else raise Not_found let rec find_projectable_vars with_evars aliases sigma y subst = let yc = normalize_alias sigma aliases y in @@ -852,8 +912,8 @@ let invert_arg_from_subst evd aliases k0 subst_in_env_extended_with_k_binders c_ let effects = ref [] in let rec aux k t = match EConstr.kind evd t with - | Rel i when i>k0+k -> aux' k (mkRel (i-k)) - | Var id -> aux' k t + | Rel i when i>k0+k -> aux' k (RelAlias (i-k)) + | Var id -> aux' k (VarAlias id) | _ -> map_with_binders evd succ aux k t and aux' k t = try project_with_effects aliases evd effects t subst_in_env_extended_with_k_binders @@ -1113,22 +1173,24 @@ let rec is_constrainable_in top evd k (ev,(fv_rels,fv_ids) as g) t = | _ -> (* We don't try to be more clever *) true let has_constrainable_free_vars env evd aliases force k ev (fv_rels,fv_ids,let_rels,let_ids) t = - let t' = expansion_of_var evd aliases t in - if t' != t then + match to_alias evd t with + | Some t -> + let expanded, t' = expansion_of_var evd aliases t in + if expanded then (* t is a local definition, we keep it only if appears in the list *) (* of let-in variables effectively occurring on the right-hand side, *) (* which is the only reason to keep it when inverting arguments *) - match EConstr.kind evd t with - | Var id -> Id.Set.mem id let_ids - | Rel n -> Int.Set.mem n let_rels - | _ -> assert false - else + match t with + | VarAlias id -> Id.Set.mem id let_ids + | RelAlias n -> Int.Set.mem n let_rels + else begin match t with + | VarAlias id -> Id.Set.mem id fv_ids + | RelAlias n -> n <= k || Int.Set.mem n fv_rels + end + | None -> (* t is an instance for a proper variable; we filter it along *) (* the free variables allowed to occur *) - match EConstr.kind evd t with - | Var id -> Id.Set.mem id fv_ids - | Rel n -> n <= k || Int.Set.mem n fv_rels - | _ -> (not force || noccur_evar env evd ev t) && is_constrainable_in true evd k (ev,(fv_rels,fv_ids)) t + (not force || noccur_evar env evd ev t) && is_constrainable_in true evd k (ev,(fv_rels,fv_ids)) t exception EvarSolvedOnTheFly of evar_map * EConstr.constr @@ -1380,12 +1442,12 @@ let rec invert_definition conv_algo choose env evd pbty (evk,argsv as ev) rhs = | (id,p)::_::_ -> if choose then (mkVar id, p) else raise (NotUniqueInType sols) in - let ty = lazy (Retyping.get_type_of env !evdref t) in + let ty = lazy (Retyping.get_type_of env !evdref (of_alias t)) in let evd = do_projection_effects (evar_define conv_algo ~choose) env ty !evdref p in evdref := evd; c with - | Not_found -> raise (NotInvertibleUsingOurAlgorithm t) + | Not_found -> raise (NotInvertibleUsingOurAlgorithm (of_alias t)) | NotUniqueInType sols -> if not !progress then raise (NotEnoughInformationToProgress sols); @@ -1396,14 +1458,14 @@ let rec invert_definition conv_algo choose env evd pbty (evk,argsv as ev) rhs = let (evd,evar,(evk',argsv' as ev')) = materialize_evar (evar_define conv_algo ~choose) env !evdref 0 ev ty' in let ts = expansions_of_var evd aliases t in - let test c = isEvar evd c || List.mem_f (EConstr.eq_constr evd) c ts in + let test c = isEvar evd c || List.exists (is_alias evd c) ts in let filter = restrict_upon_filter evd evk test argsv' in let filter = closure_of_filter evd evk' filter in let candidates = extract_candidates sols in let evd = match candidates with | NoUpdate -> let evd, ev'' = restrict_applied_evar evd ev' filter NoUpdate in - add_conv_oriented_pb ~tail:false (None,env,mkEvar ev'',t) evd + add_conv_oriented_pb ~tail:false (None,env,mkEvar ev'',of_alias t) evd | UpdateWith _ -> restrict_evar evd evk' filter candidates in @@ -1415,15 +1477,15 @@ let rec invert_definition conv_algo choose env evd pbty (evk,argsv as ev) rhs = | Rel i when i>k -> let open Context.Rel.Declaration in (match Environ.lookup_rel (i-k) env' with - | LocalAssum _ -> project_variable (mkRel (i-k)) + | LocalAssum _ -> project_variable (RelAlias (i-k)) | LocalDef (_,b,_) -> - try project_variable (mkRel (i-k)) + try project_variable (RelAlias (i-k)) with NotInvertibleUsingOurAlgorithm _ -> imitate envk (lift i (EConstr.of_constr b))) | Var id -> (match Environ.lookup_named id env' with - | LocalAssum _ -> project_variable t + | LocalAssum _ -> project_variable (VarAlias id) | LocalDef (_,b,_) -> - try project_variable t + try project_variable (VarAlias id) with NotInvertibleUsingOurAlgorithm _ -> imitate envk (EConstr.of_constr b)) | LetIn (na,b,u,c) -> imitate envk (subst1 b c) diff --git a/pretyping/evarsolve.mli b/pretyping/evarsolve.mli index f2102f8cd..b827a0ca4 100644 --- a/pretyping/evarsolve.mli +++ b/pretyping/evarsolve.mli @@ -11,6 +11,10 @@ open EConstr open Evd open Environ +type alias + +val of_alias : alias -> EConstr.t + type unification_result = | Success of evar_map | UnifFailure of evar_map * Pretype_errors.unification_error @@ -58,12 +62,12 @@ val solve_simple_eqn : conv_fun -> ?choose:bool -> env -> evar_map -> val reconsider_conv_pbs : conv_fun -> evar_map -> unification_result val is_unification_pattern_evar : env -> evar_map -> existential -> constr list -> - constr -> constr list option + constr -> alias list option val is_unification_pattern : env * int -> evar_map -> constr -> constr list -> - constr -> constr list option + constr -> alias list option -val solve_pattern_eqn : env -> evar_map -> constr list -> constr -> constr +val solve_pattern_eqn : env -> evar_map -> alias list -> constr -> constr val noccur_evar : env -> evar_map -> Evar.t -> constr -> bool diff --git a/pretyping/unification.ml b/pretyping/unification.ml index 318a0b2cd..336b3348c 100644 --- a/pretyping/unification.ml +++ b/pretyping/unification.ml @@ -206,7 +206,9 @@ let solve_pattern_eqn_array (env,nb) f l c (sigma,metasubst,evarsubst : subst0) let pb = (Conv,TypeNotProcessed) in if noccur_between sigma 1 nb c then sigma,(k,lift (-nb) c,pb)::metasubst,evarsubst - else error_cannot_unify_local env sigma (applist (f, l),c,c) + else + let l = List.map of_alias l in + error_cannot_unify_local env sigma (applist (f, l),c,c) | Evar ev -> let env' = pop_rel_context nb env in let sigma,c = pose_all_metas_as_evars env' sigma c in -- cgit v1.2.3 From 5db9588098f9f02d923c21f3914e3c671b10728f Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 24 Jan 2017 13:07:11 +0100 Subject: Quick hack to fix interpretation of patterns in Ltac. Interpretation of patterns in Ltac is essentially flawed. It does a roundtrip through the pretyper, and relies on suspicious flagging of evars in the evar source field to recognize original pattern holes. After the pattern_of_constr function was made evar-insensitive, it expanded evars that were solved by magical side-effects of the pretyper, even if it hadn't been asked to perform any heuristics. We backtrack on the insensitivity of the pattern_of_constr function. This may have a performance penalty in other dubious code, e.g. hints. In the long run we should get rid of the pattern_of_constr function. --- ltac/tacinterp.ml | 6 ++++-- plugins/quote/quote.ml | 4 ++-- pretyping/patternops.ml | 24 +++++++----------------- pretyping/patternops.mli | 2 +- tactics/hints.ml | 6 +++--- test-suite/success/change_pattern.v | 34 ++++++++++++++++++++++++++++++++++ 6 files changed, 51 insertions(+), 25 deletions(-) create mode 100644 test-suite/success/change_pattern.v diff --git a/ltac/tacinterp.ml b/ltac/tacinterp.ml index 167501de2..20ad9fd4b 100644 --- a/ltac/tacinterp.ml +++ b/ltac/tacinterp.ml @@ -691,7 +691,9 @@ let interp_pure_open_constr ist = let interp_typed_pattern ist env sigma (_,c,_) = let sigma, c = interp_gen WithoutTypeConstraint ist true pure_open_constr_flags env sigma c in - pattern_of_constr env sigma c + (** FIXME: it is necessary to be unsafe here because of the way we handle + evars in the pretyper. Sometimes they get solved eagerly. *) + pattern_of_constr env sigma (EConstr.Unsafe.to_constr c) (* Interprets a constr expression *) let interp_constr_in_compound_list inj_fun dest_fun interp_fun ist env sigma l = @@ -736,7 +738,7 @@ let interp_closed_typed_pattern_with_occurrences ist env sigma (occs, a) = try Inl (coerce_to_evaluable_ref env sigma x) with CannotCoerceTo _ -> let c = coerce_to_closed_constr env x in - Inr (pattern_of_constr env sigma c) in + Inr (pattern_of_constr env sigma (EConstr.to_constr sigma c)) in (try try_interp_ltac_var coerce_eval_ref_or_constr ist (Some (env,sigma)) (loc,id) with Not_found -> error_global_not_found ~loc (qualid_of_ident id)) diff --git a/plugins/quote/quote.ml b/plugins/quote/quote.ml index edf34607b..23069a9ab 100644 --- a/plugins/quote/quote.ml +++ b/plugins/quote/quote.ml @@ -214,9 +214,9 @@ let compute_rhs env sigma bodyi index_of_f = let i = destRel sigma (Array.last args) in PMeta (Some (coerce_meta_in i)) | App (f,args) -> - PApp (pattern_of_constr env sigma f, Array.map aux args) + PApp (pattern_of_constr env sigma (EConstr.to_constr sigma f), Array.map aux args) | Cast (c,_,_) -> aux c - | _ -> pattern_of_constr env sigma c + | _ -> pattern_of_constr env sigma (EConstr.to_constr sigma c) in aux bodyi diff --git a/pretyping/patternops.ml b/pretyping/patternops.ml index 954aa6a94..823071e29 100644 --- a/pretyping/patternops.ml +++ b/pretyping/patternops.ml @@ -122,10 +122,9 @@ let head_of_constr_reference sigma c = match EConstr.kind sigma c with | _ -> anomaly (Pp.str "Not a rigid reference") let pattern_of_constr env sigma t = - let open EConstr in let rec pattern_of_constr env t = let open Context.Rel.Declaration in - match EConstr.kind sigma t with + match kind_of_term t with | Rel n -> PRel n | Meta n -> PMeta (Some (Id.of_string ("META" ^ string_of_int n))) | Var id -> PVar id @@ -141,7 +140,7 @@ let pattern_of_constr env sigma t = pattern_of_constr (push_rel (LocalAssum (na, c)) env) b) | App (f,a) -> (match - match EConstr.kind sigma f with + match kind_of_term f with | Evar (evk,args) -> (match snd (Evd.evar_source evk sigma) with Evar_kinds.MatchingVar (true,id) -> Some id @@ -154,18 +153,14 @@ let pattern_of_constr env sigma t = | Ind (sp,u) -> PRef (canonical_gr (IndRef sp)) | Construct (sp,u) -> PRef (canonical_gr (ConstructRef sp)) | Proj (p, c) -> - pattern_of_constr env (Retyping.expand_projection env sigma p c []) + pattern_of_constr env (EConstr.to_constr sigma (Retyping.expand_projection env sigma p (EConstr.of_constr c) [])) | Evar (evk,ctxt as ev) -> (match snd (Evd.evar_source evk sigma) with | Evar_kinds.MatchingVar (b,id) -> - let ty = existential_type sigma ev in - let () = ignore (pattern_of_constr env ty) in assert (not b); PMeta (Some id) | Evar_kinds.GoalEvar -> PEvar (evk,Array.map (pattern_of_constr env) ctxt) | _ -> - let ty = existential_type sigma ev in - let () = ignore (pattern_of_constr env ty) in PMeta None) | Case (ci,p,a,br) -> let cip = @@ -179,13 +174,8 @@ let pattern_of_constr env sigma t = in PCase (cip, pattern_of_constr env p, pattern_of_constr env a, Array.to_list (Array.mapi branch_of_constr br)) - | Fix (idx, (nas, cs, ts)) -> - let inj c = EConstr.to_constr sigma c in - PFix (idx, (nas, Array.map inj cs, Array.map inj ts)) - | CoFix (idx, (nas, cs, ts)) -> - let inj c = EConstr.to_constr sigma c in - PCoFix (idx, (nas, Array.map inj cs, Array.map inj ts)) - in + | Fix f -> PFix f + | CoFix f -> PCoFix f in pattern_of_constr env t (* To process patterns, we need a translation without typing at all. *) @@ -228,7 +218,7 @@ let instantiate_pattern env sigma lvar c = ctx in let c = substl inst c in - pattern_of_constr env sigma c + pattern_of_constr env sigma (EConstr.to_constr sigma c) with Not_found (* List.index failed *) -> let vars = List.map_filter (function Name id -> Some id | _ -> None) vars in @@ -253,7 +243,7 @@ let rec subst_pattern subst pat = | PRef ref -> let ref',t = subst_global subst ref in if ref' == ref then pat else - pattern_of_constr (Global.env()) Evd.empty (EConstr.of_constr t) + pattern_of_constr (Global.env()) Evd.empty t | PVar _ | PEvar _ | PRel _ -> pat diff --git a/pretyping/patternops.mli b/pretyping/patternops.mli index 93d2c859a..5694d345c 100644 --- a/pretyping/patternops.mli +++ b/pretyping/patternops.mli @@ -39,7 +39,7 @@ val head_of_constr_reference : Evd.evar_map -> constr -> global_reference a pattern; currently, no destructor (Cases, Fix, Cofix) and no existential variable are allowed in [c] *) -val pattern_of_constr : Environ.env -> Evd.evar_map -> constr -> constr_pattern +val pattern_of_constr : Environ.env -> Evd.evar_map -> Constr.constr -> constr_pattern (** [pattern_of_glob_constr l c] translates a term [c] with metavariables into a pattern; variables bound in [l] are replaced by the pattern to which they diff --git a/tactics/hints.ml b/tactics/hints.ml index 5aacafd6f..a1c99c341 100644 --- a/tactics/hints.ml +++ b/tactics/hints.ml @@ -763,7 +763,7 @@ let make_exact_entry env sigma pri poly ?(name=PathAny) (c, cty, ctx) = match EConstr.kind sigma cty with | Prod _ -> failwith "make_exact_entry" | _ -> - let pat = Patternops.pattern_of_constr env sigma cty in + let pat = Patternops.pattern_of_constr env sigma (EConstr.to_constr sigma cty) in let hd = try head_pattern_bound pat with BoundPattern -> failwith "make_exact_entry" @@ -784,7 +784,7 @@ let make_apply_entry env sigma (eapply,hnf,verbose) pri poly ?(name=PathAny) (c, let sigma' = Evd.merge_context_set univ_flexible sigma ctx in let ce = mk_clenv_from_env env sigma' None (c,cty) in let c' = clenv_type (* ~reduce:false *) ce in - let pat = Patternops.pattern_of_constr env ce.evd c' in + let pat = Patternops.pattern_of_constr env ce.evd (EConstr.to_constr sigma c') in let hd = try head_pattern_bound pat with BoundPattern -> failwith "make_apply_entry" in @@ -934,7 +934,7 @@ let make_trivial env sigma poly ?(name=PathAny) r = let ce = mk_clenv_from_env env sigma None (c,t) in (Some hd, { pri=1; poly = poly; - pat = Some (Patternops.pattern_of_constr env ce.evd (clenv_type ce)); + pat = Some (Patternops.pattern_of_constr env ce.evd (EConstr.to_constr sigma (clenv_type ce))); name = name; db = None; secvars = secvars_of_constr env sigma c; diff --git a/test-suite/success/change_pattern.v b/test-suite/success/change_pattern.v new file mode 100644 index 000000000..874abf49f --- /dev/null +++ b/test-suite/success/change_pattern.v @@ -0,0 +1,34 @@ +Set Implicit Arguments. +Unset Strict Implicit. + +Axiom vector : Type -> nat -> Type. + +Record KleeneStore i j a := kleeneStore + { dim : nat + ; peek : vector j dim -> a + ; pos : vector i dim + }. + +Definition KSmap i j a b (f : a -> b) (s : KleeneStore i j a) : KleeneStore i j b := + kleeneStore (fun v => f (peek v)) (pos s). + +Record KleeneCoalg (i o : Type -> Type) := kleeneCoalg + { coalg :> forall a b, (o a) -> KleeneStore (i a) (i b) (o b) }. + +Axiom free_b_dim : forall i o (k : KleeneCoalg i o) a b b' (x : o a), dim (coalg k b x) = dim (coalg k b' x). +Axiom t : Type -> Type. +Axiom traverse : KleeneCoalg (fun x => x) t. + +Definition size a (x:t a) : nat := dim (traverse a a x). + +Lemma iso1_iso2_2 a (y : {x : t unit & vector a (size x)}) : False. +Proof. +destruct y. +pose (X := KSmap (traverse a unit) (traverse unit a x)). +set (e :=(eq_sym (free_b_dim traverse (a:=unit) a unit x))). +clearbody e. +(** The pattern generated by change must have holes where there were implicit + arguments in the original user-provided term. This particular example fails + if this is not the case because the inferred argument does not coincide with + the one in the considered term. *) +progress (change (dim (traverse unit a x)) with (dim X) in e). -- cgit v1.2.3 From 1523276aedcde1c79aff899ec87f05f3a708d13b Mon Sep 17 00:00:00 2001 From: Enrico Tassi Date: Thu, 9 Feb 2017 18:53:53 +0100 Subject: Missing API in EConstr. --- engine/eConstr.ml | 36 ++++++++++++++++++++++++++++++++++++ engine/eConstr.mli | 21 +++++++++++++++++++++ 2 files changed, 57 insertions(+) diff --git a/engine/eConstr.ml b/engine/eConstr.ml index 2de6d7ae4..b6b202cd9 100644 --- a/engine/eConstr.ml +++ b/engine/eConstr.ml @@ -19,12 +19,17 @@ sig type t val kind : Evd.evar_map -> t -> (t, t) Constr.kind_of_term val kind_upto : Evd.evar_map -> constr -> (constr, types) 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) 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 end = struct @@ -63,6 +68,7 @@ let rec whd_evar sigma 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 @@ -71,6 +77,11 @@ let unsafe_eq = Refl let rec to_constr sigma t = Constr.map (fun t -> to_constr sigma t) (whd_evar sigma t) +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 + end include API @@ -261,6 +272,25 @@ let decompose_lam_n_decls sigma n = 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 @@ -682,9 +712,15 @@ 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 Sigma.Sigma (t,sigma,p) = + Sigma.fresh_global ?loc ?rigid ?names env sigma reference in + Sigma.Sigma (of_constr t,sigma,p) module Unsafe = struct 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 index 7992c0633..3b479a64d 100644 --- a/engine/eConstr.mli +++ b/engine/eConstr.mli @@ -39,6 +39,8 @@ val kind_upto : Evd.evar_map -> Constr.t -> (Constr.t, Constr.t) Constr.kind_of_ 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) Constr.kind_of_term -> t @@ -143,6 +145,9 @@ 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 @@ -212,6 +217,16 @@ 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 -> + 'r Sigma.t -> Globnames.global_reference -> (t, 'r) Sigma.sigma + +(** {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 + (** {5 Unsafe operations} *) module Unsafe : @@ -219,6 +234,12 @@ 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 eq : (t, Constr.t) eq (** Use for transparent cast between types. *) end -- cgit v1.2.3 From cca57bcd89770e76e1bcc21eb41756dca2c51425 Mon Sep 17 00:00:00 2001 From: Enrico Tassi Date: Thu, 9 Feb 2017 18:53:53 +0100 Subject: Porting the ssrmatching plugin to the new EConstr API. --- plugins/ssrmatching/ssrmatching.ml4 | 62 ++++++++++++++++++++----------------- plugins/ssrmatching/ssrmatching.mli | 6 ++-- 2 files changed, 36 insertions(+), 32 deletions(-) diff --git a/plugins/ssrmatching/ssrmatching.ml4 b/plugins/ssrmatching/ssrmatching.ml4 index d4579c3a1..c40a1a9d9 100644 --- a/plugins/ssrmatching/ssrmatching.ml4 +++ b/plugins/ssrmatching/ssrmatching.ml4 @@ -301,14 +301,14 @@ let unif_EQ_args env sigma pa a = prof_unif_eq_args.profile (unif_EQ_args env sigma pa) a ;; -let unif_HO env ise p c = Evarconv.the_conv_x env (EConstr.of_constr p) (EConstr.of_constr c) ise +let unif_HO env ise p c = Evarconv.the_conv_x env p c ise -let unif_HOtype env ise p c = Evarconv.the_conv_x_leq env (EConstr.of_constr p) (EConstr.of_constr c) ise +let unif_HOtype env ise p c = Evarconv.the_conv_x_leq env p c ise let unif_HO_args env ise0 pa i ca = let n = Array.length pa in let rec loop ise j = - if j = n then ise else loop (unif_HO env ise pa.(j) ca.(i + j)) (j + 1) in + if j = n then ise else loop (unif_HO env ise (EConstr.of_constr pa.(j)) (EConstr.of_constr ca.(i + j))) (j + 1) in loop ise0 0 (* FO unification should boil down to calling w_unify with no_delta, but *) @@ -339,6 +339,7 @@ let unif_FO env ise p c = (* Perform evar substitution in main term and prune substitution. *) let nf_open_term sigma0 ise c = + let c = EConstr.Unsafe.to_constr c in let s = ise and s' = ref sigma0 in let rec nf c' = match kind_of_term c' with | Evar ex -> @@ -355,7 +356,7 @@ let nf_open_term sigma0 ise c = | Evar_defined c' -> s' := Evd.define k (nf c') !s' | _ -> () in let c' = nf c in let _ = Evd.fold copy_def sigma0 () in - !s', Evd.evar_universe_context s, c' + !s', Evd.evar_universe_context s, EConstr.of_constr c' let unif_end env sigma0 ise0 pt ok = let ise = Evarconv.consider_remaining_unif_problems env ise0 in @@ -429,7 +430,7 @@ type tpattern = { up_a : constr array; up_t : constr; (* equation proof term or matched term *) up_dir : ssrdir; (* direction of the rule *) - up_ok : constr -> evar_map -> bool; (* progess test for rewrite *) + up_ok : constr -> evar_map -> bool; (* progress test for rewrite *) } let all_ok _ _ = true @@ -632,13 +633,14 @@ let match_upats_FO upats env sigma0 ise orig_c = | _ -> unif_FO env ise u.up_FO c' in let ise' = (* Unify again using HO to assign evars *) let p = mkApp (u.up_f, u.up_a) in - try unif_HO env ise p c' with _ -> raise NoMatch in + try unif_HO env ise (EConstr.of_constr p) (EConstr.of_constr c') with e when CErrors.noncritical e -> raise NoMatch in let lhs = mkSubApp f i a in - let pt' = unif_end env sigma0 ise' u.up_t (u.up_ok lhs) in + let pt' = unif_end env sigma0 ise' (EConstr.of_constr u.up_t) (u.up_ok lhs) in + let pt' = pi1 pt', pi2 pt', EConstr.Unsafe.to_constr (pi3 pt') in raise (FoundUnif (ungen_upat lhs pt' u)) with FoundUnif (s,_,_) as sig_u when dont_impact_evars s -> raise sig_u | Not_found -> CErrors.anomaly (str"incomplete ise in match_upats_FO") - | _ -> () in + | e when CErrors.noncritical e -> () in List.iter one_match fpats done; iter_constr_LR loop f; Array.iter loop a in @@ -651,7 +653,7 @@ let match_upats_FO upats env sigma0 ise c = let match_upats_HO ~on_instance upats env sigma0 ise c = - let dont_impact_evars = dont_impact_evars_in c in + let dont_impact_evars = dont_impact_evars_in c in let it_did_match = ref false in let failed_because_of_TC = ref false in let rec aux upats env sigma0 ise c = @@ -673,16 +675,17 @@ let match_upats_HO ~on_instance upats env sigma0 ise c = | KpatLet -> let x, v, t, b = destLetIn f in let _, pv, _, pb = destLetIn u.up_f in - let ise' = unif_HO env ise pv v in + let ise' = unif_HO env ise (EConstr.of_constr pv) (EConstr.of_constr v) in unif_HO (Environ.push_rel (Context.Rel.Declaration.LocalAssum(x, t)) env) - ise' pb b + ise' (EConstr.of_constr pb) (EConstr.of_constr b) | KpatFlex | KpatProj _ -> - unif_HO env ise u.up_f (mkSubApp f (i - Array.length u.up_a) a) - | _ -> unif_HO env ise u.up_f f in + unif_HO env ise (EConstr.of_constr u.up_f) (EConstr.of_constr(mkSubApp f (i - Array.length u.up_a) a)) + | _ -> unif_HO env ise (EConstr.of_constr u.up_f) (EConstr.of_constr f) in let ise'' = unif_HO_args env ise' u.up_a (i - Array.length u.up_a) a in let lhs = mkSubApp f i a in - let pt' = unif_end env sigma0 ise'' u.up_t (u.up_ok lhs) in + let pt' = unif_end env sigma0 ise'' (EConstr.of_constr u.up_t) (u.up_ok lhs) in + let pt' = pi1 pt', pi2 pt', EConstr.Unsafe.to_constr (pi3 pt') in on_instance (ungen_upat lhs pt' u) with FoundUnif (s,_,_) as sig_u when dont_impact_evars s -> raise sig_u | NoProgress -> it_did_match := true @@ -721,13 +724,13 @@ let assert_done_multires r = r := Some (n+1,xs); try List.nth xs n with Failure _ -> raise NoMatch -type subst = Environ.env -> Term.constr -> Term.constr -> int -> Term.constr +type subst = Environ.env -> constr -> constr -> int -> constr type find_P = - Environ.env -> Term.constr -> int -> + Environ.env -> constr -> int -> k:subst -> - Term.constr + constr type conclude = unit -> - Term.constr * ssrdir * (Evd.evar_map * Evd.evar_universe_context * Term.constr) + constr * ssrdir * (Evd.evar_map * Evd.evar_universe_context * constr) (* upats_origin makes a better error message only *) let mk_tpattern_matcher ?(all_instances=false) @@ -897,7 +900,7 @@ let pr_pattern_aux pr_constr = function | E_As_X_In_T (e,x,t) -> pr_constr e ++ str " as " ++ pr_constr x ++ str " in " ++ pr_constr t let pp_pattern (sigma, p) = - pr_pattern_aux (fun t -> pr_constr_pat (pi3 (nf_open_term sigma sigma t))) p + pr_pattern_aux (fun t -> pr_constr_pat (EConstr.Unsafe.to_constr (pi3 (nf_open_term sigma sigma (EConstr.of_constr t))))) p let pr_cpattern = pr_term let pr_rpattern _ _ _ = pr_pattern @@ -1001,7 +1004,7 @@ type occ = (bool * int list) option type rpattern = (cpattern, cpattern) ssrpattern let pr_rpattern = pr_pattern -type pattern = Evd.evar_map * (Term.constr, Term.constr) ssrpattern +type pattern = Evd.evar_map * (constr, constr) ssrpattern let id_of_cpattern = function @@ -1260,7 +1263,7 @@ let eval_pattern ?raise_NoMatch env0 sigma0 concl0 pattern occ do_subst = let holep = mk_upat_for env0 sigma (sigma, hole) all_ok in let find_X, end_X = mk_tpattern_matcher ?raise_NoMatch sigma occ holep in let concl = find_T env0 concl0 1 (fun env c _ h -> - let p_sigma = unify_HO env (create_evar_defs sigma) c p in + let p_sigma = unify_HO env (create_evar_defs sigma) (EConstr.of_constr c) (EConstr.of_constr p) in let sigma, e_body = pop_evar p_sigma ex p in fs p_sigma (find_X env (fs sigma p) h (fun env _ -> do_subst env e_body))) in @@ -1276,7 +1279,7 @@ let eval_pattern ?raise_NoMatch env0 sigma0 concl0 pattern occ do_subst = let re = mk_upat_for env0 sigma0 (sigma, e) all_ok in let find_E, end_E = mk_tpattern_matcher ?raise_NoMatch sigma0 occ re in let concl = find_T env0 concl0 1 (fun env c _ h -> - let p_sigma = unify_HO env (create_evar_defs sigma) c p in + let p_sigma = unify_HO env (create_evar_defs sigma) (EConstr.of_constr c) (EConstr.of_constr p) in let sigma, e_body = pop_evar p_sigma ex p in fs p_sigma (find_X env (fs sigma p) h (fun env c _ h -> find_E env e_body h do_subst))) in @@ -1286,17 +1289,17 @@ let eval_pattern ?raise_NoMatch env0 sigma0 concl0 pattern occ do_subst = let p, e = fs sigma p, fs sigma e in let ex = ex_value hole in let rp = - let e_sigma = unify_HO env0 sigma hole e in + let e_sigma = unify_HO env0 sigma (EConstr.of_constr hole) (EConstr.of_constr e) in e_sigma, fs e_sigma p in let rp = mk_upat_for ~hack:true env0 sigma0 rp all_ok in let find_TE, end_TE = mk_tpattern_matcher sigma0 noindex rp in let holep = mk_upat_for env0 sigma (sigma, hole) all_ok in let find_X, end_X = mk_tpattern_matcher sigma occ holep in let concl = find_TE env0 concl0 1 (fun env c _ h -> - let p_sigma = unify_HO env (create_evar_defs sigma) c p in + let p_sigma = unify_HO env (create_evar_defs sigma) (EConstr.of_constr c) (EConstr.of_constr p) in let sigma, e_body = pop_evar p_sigma ex p in fs p_sigma (find_X env (fs sigma p) h (fun env c _ h -> - let e_sigma = unify_HO env sigma e_body e in + let e_sigma = unify_HO env sigma (EConstr.of_constr e_body) (EConstr.of_constr e) in let e_body = fs e_sigma e in do_subst env e_body e_body h))) in let _ = end_X () in let _ = end_TE () in @@ -1332,13 +1335,15 @@ let mk_tpattern ?p_origin env sigma0 sigma_t f dir c = ;; let pf_fill_occ env concl occ sigma0 p (sigma, t) ok h = + let p = EConstr.Unsafe.to_constr p in + let concl = EConstr.Unsafe.to_constr concl in let ise = create_evar_defs sigma in - let ise, u = mk_tpattern env sigma0 (ise,t) ok L2R p in + let ise, u = mk_tpattern env sigma0 (ise,EConstr.Unsafe.to_constr t) ok L2R p in let find_U, end_U = mk_tpattern_matcher ~raise_NoMatch:true sigma0 occ (ise,[u]) in let concl = find_U env concl h (fun _ _ _ -> mkRel) in let rdx, _, (sigma, uc, p) = end_U () in - sigma, uc, p, concl, rdx + sigma, uc, EConstr.of_constr p, EConstr.of_constr concl, EConstr.of_constr rdx let fill_occ_term env cl occ sigma0 (sigma, t) = try @@ -1351,12 +1356,11 @@ let fill_occ_term env cl occ sigma0 (sigma, t) = if sigma' != sigma0 then raise NoMatch else cl, (Evd.merge_universe_context sigma' uc, t') with _ -> - errorstrm (str "partial term " ++ pr_constr_pat t + errorstrm (str "partial term " ++ pr_constr_pat (EConstr.Unsafe.to_constr t) ++ str " does not match any subterm of the goal") let pf_fill_occ_term gl occ t = let sigma0 = project gl and env = pf_env gl and concl = pf_concl gl in - let concl = EConstr.Unsafe.to_constr concl in let cl,(_,t) = fill_occ_term env concl occ sigma0 t in cl, t diff --git a/plugins/ssrmatching/ssrmatching.mli b/plugins/ssrmatching/ssrmatching.mli index 74a603e51..fa0c2f5b1 100644 --- a/plugins/ssrmatching/ssrmatching.mli +++ b/plugins/ssrmatching/ssrmatching.mli @@ -194,7 +194,7 @@ val mk_tpattern_matcher : (* convenience shortcut: [pf_fill_occ_term gl occ (sigma,t)] returns * the conclusion of [gl] where [occ] occurrences of [t] have been replaced * by [Rel 1] and the instance of [t] *) -val pf_fill_occ_term : goal sigma -> occ -> evar_map * constr -> constr * constr +val pf_fill_occ_term : goal sigma -> occ -> evar_map * EConstr.t -> EConstr.t * EConstr.t (* It may be handy to inject a simple term into the first form of cpattern *) val cpattern_of_term : char * glob_constr_and_expr -> cpattern @@ -216,8 +216,8 @@ val assert_done : 'a option ref -> 'a [consider_remaining_unif_problems] and [resolve_typeclasses]. In case of failure they raise [NoMatch] *) -val unify_HO : env -> evar_map -> constr -> constr -> evar_map -val pf_unify_HO : goal sigma -> constr -> constr -> goal sigma +val unify_HO : env -> evar_map -> EConstr.constr -> EConstr.constr -> evar_map +val pf_unify_HO : goal sigma -> EConstr.constr -> EConstr.constr -> goal sigma (** Some more low level functions needed to implement the full SSR language on top of the former APIs *) -- cgit v1.2.3 From 486acdd7b50d4fdc0956011b7b48dc6ba96dd4a8 Mon Sep 17 00:00:00 2001 From: Maxime Dénès Date: Fri, 24 Mar 2017 23:45:54 +0100 Subject: Fix interpretation of Ltac patterns episode 2. After 5db9588098f9f, some extra evar-normalization remained (compared to trunk) that would change the semantics e.g. of change bindings under Ltac match. This is just circumventing a fundamental flaw in the treatment of patterns. --- pretyping/patternops.ml | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/pretyping/patternops.ml b/pretyping/patternops.ml index 8c6b39b7e..d6a7c5192 100644 --- a/pretyping/patternops.ml +++ b/pretyping/patternops.ml @@ -153,7 +153,7 @@ let pattern_of_constr env sigma t = | Ind (sp,u) -> PRef (canonical_gr (IndRef sp)) | Construct (sp,u) -> PRef (canonical_gr (ConstructRef sp)) | Proj (p, c) -> - pattern_of_constr env (EConstr.to_constr sigma (Retyping.expand_projection env sigma p (EConstr.of_constr c) [])) + pattern_of_constr env (EConstr.Unsafe.to_constr (Retyping.expand_projection env sigma p (EConstr.of_constr c) [])) | Evar (evk,ctxt as ev) -> (match snd (Evd.evar_source evk sigma) with | Evar_kinds.MatchingVar (b,id) -> @@ -218,7 +218,7 @@ let instantiate_pattern env sigma lvar c = ctx in let c = substl inst c in - pattern_of_constr env sigma (EConstr.to_constr sigma c) + pattern_of_constr env sigma (EConstr.Unsafe.to_constr c) with Not_found (* List.index failed *) -> let vars = List.map_filter (function Name id -> Some id | _ -> None) vars in -- cgit v1.2.3 From 98e51078fa624ce780b16d8e372ef33ac97ffaee Mon Sep 17 00:00:00 2001 From: Maxime Dénès Date: Sat, 25 Mar 2017 00:00:18 +0100 Subject: Mathcomp overlay. --- dev/ci/ci-user-overlay.sh | 9 +++++++++ 1 file changed, 9 insertions(+) diff --git a/dev/ci/ci-user-overlay.sh b/dev/ci/ci-user-overlay.sh index 028574743..ab2d2e382 100644 --- a/dev/ci/ci-user-overlay.sh +++ b/dev/ci/ci-user-overlay.sh @@ -20,3 +20,12 @@ # the name of the branch from which the PR originated. "" if the # current job is a push build. +echo $TRAVIS_PULL_REQUEST_BRANCH +echo $TRAVIS_PULL_REQUEST +echo $TRAVIS_BRANCH +echo $TRAVIS_COMMIT + +if [ $TRAVIS_PULL_REQUEST == "379" ] || [ $TRAVIS_BRANCH == "pr379" ]; then + mathcomp_CI_BRANCH=econstr + mathcomp_CI_GITURL=https://github.com/maximedenes/math-comp.git +fi -- cgit v1.2.3 From dc8d8daf8850ff1a414ae36c860bc925d87eab01 Mon Sep 17 00:00:00 2001 From: Maxime Dénès Date: Tue, 28 Mar 2017 18:15:02 +0200 Subject: Revert to incorrect heuristic in apply. Was breaking e.g. fiat-crypto. --- pretyping/unification.ml | 14 +++++++++++++- 1 file changed, 13 insertions(+), 1 deletion(-) diff --git a/pretyping/unification.ml b/pretyping/unification.ml index 8824c06ab..611d165fe 100644 --- a/pretyping/unification.ml +++ b/pretyping/unification.ml @@ -64,6 +64,18 @@ let _ = Goptions.declare_bool_option { Goptions.optwrite = (fun a -> debug_unification:=a); } +(** Making this unification algorithm correct w.r.t. the evar-map abstraction + breaks too much stuff. So we redefine incorrect functions here. *) + +let unsafe_occur_meta_or_existential c = + let c = EConstr.Unsafe.to_constr c in + let rec occrec c = match kind_of_term c with + | Evar _ -> raise Occur + | Meta _ -> raise Occur + | _ -> iter_constr occrec c + in try occrec c; false with Occur -> true + + let occur_meta_or_undefined_evar evd c = (** This is performance-critical. Using the evar-insensitive API changes the resulting heuristic. *) @@ -1880,7 +1892,7 @@ let w_unify_to_subterm_list env evd flags hdmeta oplist t = else let allow_K = flags.allow_K_in_toplevel_higher_order_unification in let flags = - if occur_meta_or_existential evd op || !keyed_unification then + if unsafe_occur_meta_or_existential op || !keyed_unification then (* This is up to delta for subterms w/o metas ... *) flags else -- cgit v1.2.3 From 1cf91df6e0617c316dff7d99c7603a26b694e647 Mon Sep 17 00:00:00 2001 From: Maxime Dénès Date: Wed, 29 Mar 2017 16:04:15 +0200 Subject: Fix call to broken unsafe_type_of in apply tactic. This broke the build of iris-coq in the EConstr branch. Each time you use unsafe_type_of, I loose a night of sleep, so please stop. --- proofs/logic.ml | 11 +++++++---- 1 file changed, 7 insertions(+), 4 deletions(-) diff --git a/proofs/logic.ml b/proofs/logic.ml index a31cadd88..b38a901c2 100644 --- a/proofs/logic.ml +++ b/proofs/logic.ml @@ -334,8 +334,10 @@ let meta_free_prefix sigma a = with Stop acc -> Array.rev_of_list acc let goal_type_of env sigma c = - if !check then EConstr.Unsafe.to_constr (unsafe_type_of env sigma (EConstr.of_constr c)) - else EConstr.Unsafe.to_constr (Retyping.get_type_of env sigma (EConstr.of_constr c)) + if !check then + let (sigma,t) = type_of env sigma (EConstr.of_constr c) in + (sigma, EConstr.Unsafe.to_constr t) + else (sigma, EConstr.Unsafe.to_constr (Retyping.get_type_of env sigma (EConstr.of_constr c))) let rec mk_refgoals sigma goal goalacc conclty trm = let env = Goal.V82.env sigma goal in @@ -418,7 +420,7 @@ let rec mk_refgoals sigma goal goalacc conclty trm = | _ -> if occur_meta sigma (EConstr.of_constr trm) then anomaly (Pp.str "refiner called with a meta in non app/case subterm"); - let t'ty = goal_type_of env sigma trm in + let (sigma, t'ty) = goal_type_of env sigma trm in let sigma = check_conv_leq_goal env sigma trm t'ty conclty in (goalacc,t'ty,sigma, trm) @@ -478,7 +480,8 @@ and mk_hdgoals sigma goal goalacc trm = | _ -> if !check && occur_meta sigma (EConstr.of_constr trm) then anomaly (Pp.str "refine called with a dependent meta"); - goalacc, goal_type_of env sigma trm, sigma, trm + let (sigma, ty) = goal_type_of env sigma trm in + goalacc, ty, sigma, trm and mk_arggoals sigma goal goalacc funty allargs = let foldmap (goalacc, funty, sigma) harg = -- cgit v1.2.3 From b5f07be9fdcd41fdaf73503e5214e766bf6a303b Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Thu, 30 Mar 2017 16:22:33 +0200 Subject: Specially crafted EConstr.kind to be more efficient. We do one step of loop unrolling, limit the number of allocations and mark the function as inline. --- engine/eConstr.ml | 56 ++++++++++++++++++++++++++++++++++++++++++++++++------- 1 file changed, 49 insertions(+), 7 deletions(-) diff --git a/engine/eConstr.ml b/engine/eConstr.ml index f50a8b850..68ef37226 100644 --- a/engine/eConstr.ml +++ b/engine/eConstr.ml @@ -41,8 +41,8 @@ let safe_evar_value sigma ev = let rec whd_evar sigma c = match Constr.kind c with - | Evar (evk, args) -> - (match safe_evar_value sigma (evk, args) with + | Evar ev -> + (match safe_evar_value sigma ev with Some c -> whd_evar sigma c | None -> c) | Sort (Type u) -> @@ -57,16 +57,58 @@ let rec whd_evar sigma c = | 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') - | App (f, args) when isEvar f -> + | App (f, args) -> (** Enforce smart constructor invariant on applications *) - let ev = destEvar f in + begin match Constr.kind f with + | Evar ev -> + begin match safe_evar_value sigma ev with + | None -> c + | Some f -> whd_evar sigma (mkApp (f, args)) + end + | _ -> c + end + | _ -> c + +(** Loop unrolling to ensure efficiency *) +let kind sigma c = + let c = Constr.kind c in + match c with + | Evar ev -> begin match safe_evar_value sigma ev with + | Some c -> Constr.kind (whd_evar sigma c) | None -> c - | Some f -> whd_evar sigma (mkApp (f, args)) end - | _ -> c + | Sort (Type u) -> + let u' = Evd.normalize_universe sigma u in + if u' == u then c else Sort (Sorts.sort_of_univ u') + | Const (c', u) -> + if Univ.Instance.is_empty u then c + else + let u' = Evd.normalize_universe_instance sigma u in + if u' == u then c else Const (c', u') + | Ind (i, u) -> + if Univ.Instance.is_empty u then c + else + let u' = Evd.normalize_universe_instance sigma u in + if u' == u then c else Ind (i, u') + | Construct (co, u) -> + if Univ.Instance.is_empty u then c + else + let u' = Evd.normalize_universe_instance sigma u in + if u' == u then c else Construct (co, u') + | App (f, args) -> + (** Enforce smart constructor invariant on applications *) + begin match Constr.kind f with + | Evar ev -> + begin match safe_evar_value sigma ev with + | None -> c + | Some f -> Constr.kind (whd_evar sigma (mkApp (f, args))) + end + | _ -> c + end + | c -> c +[@@ocaml.inline] -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 -- cgit v1.2.3 From a20494163815e4b8275c4d0412d6c857c95263f4 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 31 Mar 2017 22:39:36 +0200 Subject: Revert "Specially crafted EConstr.kind to be more efficient." This reverts commit b5f07be9fdcd41fdaf73503e5214e766bf6a303b. The performance difference was not conclusive enough to pay for the code ugliness. --- engine/eConstr.ml | 56 +++++++------------------------------------------------ 1 file changed, 7 insertions(+), 49 deletions(-) diff --git a/engine/eConstr.ml b/engine/eConstr.ml index 68ef37226..f50a8b850 100644 --- a/engine/eConstr.ml +++ b/engine/eConstr.ml @@ -41,8 +41,8 @@ let safe_evar_value sigma ev = let rec whd_evar sigma c = match Constr.kind c with - | Evar ev -> - (match safe_evar_value sigma ev with + | Evar (evk, args) -> + (match safe_evar_value sigma (evk, args) with Some c -> whd_evar sigma c | None -> c) | Sort (Type u) -> @@ -57,58 +57,16 @@ let rec whd_evar sigma c = | 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') - | App (f, args) -> + | App (f, args) when isEvar f -> (** Enforce smart constructor invariant on applications *) - begin match Constr.kind f with - | Evar ev -> - begin match safe_evar_value sigma ev with - | None -> c - | Some f -> whd_evar sigma (mkApp (f, args)) - end - | _ -> c - end - | _ -> c - -(** Loop unrolling to ensure efficiency *) -let kind sigma c = - let c = Constr.kind c in - match c with - | Evar ev -> + let ev = destEvar f in begin match safe_evar_value sigma ev with - | Some c -> Constr.kind (whd_evar sigma c) | None -> c + | Some f -> whd_evar sigma (mkApp (f, args)) end - | Sort (Type u) -> - let u' = Evd.normalize_universe sigma u in - if u' == u then c else Sort (Sorts.sort_of_univ u') - | Const (c', u) -> - if Univ.Instance.is_empty u then c - else - let u' = Evd.normalize_universe_instance sigma u in - if u' == u then c else Const (c', u') - | Ind (i, u) -> - if Univ.Instance.is_empty u then c - else - let u' = Evd.normalize_universe_instance sigma u in - if u' == u then c else Ind (i, u') - | Construct (co, u) -> - if Univ.Instance.is_empty u then c - else - let u' = Evd.normalize_universe_instance sigma u in - if u' == u then c else Construct (co, u') - | App (f, args) -> - (** Enforce smart constructor invariant on applications *) - begin match Constr.kind f with - | Evar ev -> - begin match safe_evar_value sigma ev with - | None -> c - | Some f -> Constr.kind (whd_evar sigma (mkApp (f, args))) - end - | _ -> c - end - | c -> c -[@@ocaml.inline] + | _ -> 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 -- cgit v1.2.3 From 6add354ad9ca0f68d3ef311c4e53ee96d9fdb4d7 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 31 Mar 2017 22:46:05 +0200 Subject: Ensuring proper cast invariants in EConstr.kind. The kernel does fishy things with casts, such that ensuring there are no two consecutive VMcast or NATIVEcast in terms. We enforce this in EConstr view as well. --- engine/eConstr.ml | 16 ++++++++++++---- 1 file changed, 12 insertions(+), 4 deletions(-) diff --git a/engine/eConstr.ml b/engine/eConstr.ml index f50a8b850..9026800f2 100644 --- a/engine/eConstr.ml +++ b/engine/eConstr.ml @@ -41,10 +41,11 @@ let safe_evar_value sigma ev = let rec whd_evar sigma c = match Constr.kind c with - | Evar (evk, args) -> - (match safe_evar_value sigma (evk, args) with - Some c -> whd_evar sigma c - | None -> c) + | Evar ev -> + begin match safe_evar_value sigma ev with + | Some c -> whd_evar sigma c + | None -> 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') @@ -64,6 +65,13 @@ let rec whd_evar sigma c = | 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) -- cgit v1.2.3 From ce029533a1f0fc6ac9e28d162350a64446522246 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 31 Mar 2017 23:05:17 +0200 Subject: Make the Constr.kind_of_term type parametric in sorts and universes. --- engine/eConstr.ml | 6 +++--- engine/eConstr.mli | 6 +++--- engine/evarutil.mli | 3 ++- engine/universes.mli | 4 ++-- kernel/constr.ml | 12 ++++++------ kernel/constr.mli | 22 +++++++++++----------- kernel/term.ml | 11 ++++++----- kernel/term.mli | 13 +++++++------ pretyping/reductionops.mli | 2 +- 9 files changed, 41 insertions(+), 38 deletions(-) diff --git a/engine/eConstr.ml b/engine/eConstr.ml index 9026800f2..094841d69 100644 --- a/engine/eConstr.ml +++ b/engine/eConstr.ml @@ -17,11 +17,11 @@ open Evd module API : sig type t -val kind : Evd.evar_map -> t -> (t, t) Constr.kind_of_term -val kind_upto : Evd.evar_map -> constr -> (constr, types) Constr.kind_of_term +val kind : Evd.evar_map -> t -> (t, t, Sorts.t, Univ.Instance.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) Constr.kind_of_term -> t +val of_kind : (t, t, Sorts.t, Univ.Instance.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 diff --git a/engine/eConstr.mli b/engine/eConstr.mli index 1ae71216f..e7287fc06 100644 --- a/engine/eConstr.mli +++ b/engine/eConstr.mli @@ -30,11 +30,11 @@ type rel_context = (constr, types) Context.Rel.pt (** {5 Destructors} *) -val kind : Evd.evar_map -> t -> (t, t) Constr.kind_of_term +val kind : Evd.evar_map -> t -> (t, t, Sorts.t, Univ.Instance.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) Constr.kind_of_term +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}. *) @@ -43,7 +43,7 @@ val kind_of_type : Evd.evar_map -> t -> (t, t) Term.kind_of_type (** {5 Constructors} *) -val of_kind : (t, t) Constr.kind_of_term -> t +val of_kind : (t, t, Sorts.t, Univ.Instance.t) Constr.kind_of_term -> t (** Construct a term from a view. *) val of_constr : Constr.t -> t diff --git a/engine/evarutil.mli b/engine/evarutil.mli index da49d4e11..ca9591e71 100644 --- a/engine/evarutil.mli +++ b/engine/evarutil.mli @@ -178,7 +178,8 @@ val flush_and_check_evars : evar_map -> constr -> Constr.constr (** 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 -> (Constr.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 diff --git a/engine/universes.mli b/engine/universes.mli index c3e2055f3..932de941a 100644 --- a/engine/universes.mli +++ b/engine/universes.mli @@ -76,8 +76,8 @@ val eq_constr_univs_infer : UGraph.t -> 'a constraint_accumulator -> {!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) kind_of_term) -> - (constr -> (constr,types) kind_of_term) -> + (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] diff --git a/kernel/constr.ml b/kernel/constr.ml index 0ae3fb474..5a7561bf5 100644 --- a/kernel/constr.ml +++ b/kernel/constr.ml @@ -81,27 +81,27 @@ type pconstructor = constructor puniverses (* [Var] is used for named variables and [Rel] for variables as de Bruijn indices. *) -type ('constr, 'types) kind_of_term = +type ('constr, 'types, 'sort, 'univs) kind_of_term = | Rel of int | Var of Id.t | Meta of metavariable | Evar of 'constr pexistential - | Sort of Sorts.t + | Sort of 'sort | Cast of 'constr * cast_kind * 'types | Prod of Name.t * 'types * 'types | Lambda of Name.t * 'types * 'constr | LetIn of Name.t * 'constr * 'types * 'constr | App of 'constr * 'constr array - | Const of pconstant - | Ind of pinductive - | Construct of pconstructor + | Const of (constant * 'univs) + | Ind of (inductive * 'univs) + | Construct of (constructor * 'univs) | Case of case_info * 'constr * 'constr * 'constr array | Fix of ('constr, 'types) pfixpoint | CoFix of ('constr, 'types) pcofixpoint | Proj of projection * 'constr (* constr is the fixpoint of the previous type. Requires option -rectypes of the Caml compiler to be set *) -type t = (t,t) kind_of_term +type t = (t, t, Sorts.t, Instance.t) kind_of_term type constr = t type existential = existential_key * constr array diff --git a/kernel/constr.mli b/kernel/constr.mli index c8be89fe2..700c235e6 100644 --- a/kernel/constr.mli +++ b/kernel/constr.mli @@ -188,7 +188,7 @@ type ('constr, 'types) pfixpoint = type ('constr, 'types) pcofixpoint = int * ('constr, 'types) prec_declaration -type ('constr, 'types) kind_of_term = +type ('constr, 'types, 'sort, 'univs) kind_of_term = | Rel of int (** Gallina-variable introduced by [forall], [fun], [let-in], [fix], or [cofix]. *) | Var of Id.t (** Gallina-variable that was introduced by Vernacular-command that extends @@ -197,7 +197,7 @@ type ('constr, 'types) kind_of_term = | Meta of metavariable | Evar of 'constr pexistential - | Sort of Sorts.t + | Sort of 'sort | Cast of 'constr * cast_kind * 'types | Prod of Name.t * 'types * 'types (** Concrete syntax ["forall A:B,C"] is represented as [Prod (A,B,C)]. *) | Lambda of Name.t * 'types * 'constr (** Concrete syntax ["fun A:B => C"] is represented as [Lambda (A,B,C)]. *) @@ -208,11 +208,11 @@ type ('constr, 'types) kind_of_term = - [F] itself is not {!App} - and [[|P1;..;Pn|]] is not empty. *) - | Const of constant puniverses (** Gallina-variable that was introduced by Vernacular-command that extends the global environment + | Const of (constant * 'univs) (** Gallina-variable that was introduced by Vernacular-command that extends the global environment (i.e. [Parameter], or [Axiom], or [Definition], or [Theorem] etc.) *) - | Ind of inductive puniverses (** A name of an inductive type defined by [Variant], [Inductive] or [Record] Vernacular-commands. *) - | Construct of constructor puniverses (** A constructor of an inductive type defined by [Variant], [Inductive] or [Record] Vernacular-commands. *) + | Ind of (inductive * 'univs) (** A name of an inductive type defined by [Variant], [Inductive] or [Record] Vernacular-commands. *) + | Construct of (constructor * 'univs) (** A constructor of an inductive type defined by [Variant], [Inductive] or [Record] Vernacular-commands. *) | Case of case_info * 'constr * 'constr * 'constr array | Fix of ('constr, 'types) pfixpoint | CoFix of ('constr, 'types) pcofixpoint @@ -222,8 +222,8 @@ type ('constr, 'types) kind_of_term = least one argument and the function is not itself an applicative term *) -val kind : constr -> (constr, types) kind_of_term -val of_kind : (constr, types) kind_of_term -> constr +val kind : constr -> (constr, types, Sorts.t, Univ.Instance.t) kind_of_term +val of_kind : (constr, types, Sorts.t, Univ.Instance.t) kind_of_term -> constr (** [equal a b] is true if [a] equals [b] modulo alpha, casts, and application grouping *) @@ -312,8 +312,8 @@ val compare_head_gen : (bool -> Univ.Instance.t -> Univ.Instance.t -> bool) -> constr -> constr -> bool val compare_head_gen_leq_with : - (constr -> (constr,types) kind_of_term) -> - (constr -> (constr,types) kind_of_term) -> + (constr -> (constr, types, Sorts.t, Univ.Instance.t) kind_of_term) -> + (constr -> (constr, types, Sorts.t, Univ.Instance.t) kind_of_term) -> (bool -> Univ.Instance.t -> Univ.Instance.t -> bool) -> (Sorts.t -> Sorts.t -> bool) -> (constr -> constr -> bool) -> @@ -325,8 +325,8 @@ val compare_head_gen_leq_with : is used,rather than {!kind}, to expose the immediate subterms of [c1] (resp. [c2]). *) val compare_head_gen_with : - (constr -> (constr,types) kind_of_term) -> - (constr -> (constr,types) kind_of_term) -> + (constr -> (constr, types, Sorts.t, Univ.Instance.t) kind_of_term) -> + (constr -> (constr, types, Sorts.t, Univ.Instance.t) kind_of_term) -> (bool -> Univ.Instance.t -> Univ.Instance.t -> bool) -> (Sorts.t -> Sorts.t -> bool) -> (constr -> constr -> bool) -> diff --git a/kernel/term.ml b/kernel/term.ml index 3881cd12d..e5a681375 100644 --- a/kernel/term.ml +++ b/kernel/term.ml @@ -71,20 +71,21 @@ type pconstant = constant puniverses type pinductive = inductive puniverses type pconstructor = constructor puniverses -type ('constr, 'types) kind_of_term = ('constr, 'types) Constr.kind_of_term = +type ('constr, 'types, 'sort, 'univs) kind_of_term = + ('constr, 'types, 'sort, 'univs) Constr.kind_of_term = | Rel of int | Var of Id.t | Meta of metavariable | Evar of 'constr pexistential - | Sort of sorts + | Sort of 'sort | Cast of 'constr * cast_kind * 'types | Prod of Name.t * 'types * 'types | Lambda of Name.t * 'types * 'constr | LetIn of Name.t * 'constr * 'types * 'constr | App of 'constr * 'constr array - | Const of pconstant - | Ind of pinductive - | Construct of pconstructor + | Const of (constant * 'univs) + | Ind of (inductive * 'univs) + | Construct of (constructor * 'univs) | Case of case_info * 'constr * 'constr * 'constr array | Fix of ('constr, 'types) pfixpoint | CoFix of ('constr, 'types) pcofixpoint diff --git a/kernel/term.mli b/kernel/term.mli index e393adb81..a9bb67705 100644 --- a/kernel/term.mli +++ b/kernel/term.mli @@ -68,20 +68,21 @@ type ('constr, 'types) prec_declaration = type ('constr, 'types) pfixpoint = ('constr, 'types) Constr.pfixpoint type ('constr, 'types) pcofixpoint = ('constr, 'types) Constr.pcofixpoint -type ('constr, 'types) kind_of_term = ('constr, 'types) Constr.kind_of_term = +type ('constr, 'types, 'sort, 'univs) kind_of_term = + ('constr, 'types, 'sort, 'univs) Constr.kind_of_term = | Rel of int | Var of Id.t | Meta of metavariable | Evar of 'constr pexistential - | Sort of sorts + | Sort of 'sort | Cast of 'constr * cast_kind * 'types | Prod of Name.t * 'types * 'types | Lambda of Name.t * 'types * 'constr | LetIn of Name.t * 'constr * 'types * 'constr | App of 'constr * 'constr array - | Const of constant puniverses - | Ind of inductive puniverses - | Construct of constructor puniverses + | Const of (constant * 'univs) + | Ind of (inductive * 'univs) + | Construct of (constructor * 'univs) | Case of case_info * 'constr * 'constr * 'constr array | Fix of ('constr, 'types) pfixpoint | CoFix of ('constr, 'types) pcofixpoint @@ -443,7 +444,7 @@ val leq_constr_univs : constr UGraph.check_function application grouping and ignoring universe instances. *) val eq_constr_nounivs : constr -> constr -> bool -val kind_of_term : constr -> (constr, types) kind_of_term +val kind_of_term : constr -> (constr, types, Sorts.t, Univ.Instance.t) kind_of_term (** Alias for [Constr.kind] *) val constr_ord : constr -> constr -> int diff --git a/pretyping/reductionops.mli b/pretyping/reductionops.mli index 15ddeb15c..18416b142 100644 --- a/pretyping/reductionops.mli +++ b/pretyping/reductionops.mli @@ -232,7 +232,7 @@ type 'a miota_args = { val reducible_mind_case : evar_map -> constr -> bool val reduce_mind_case : evar_map -> constr miota_args -> constr -val find_conclusion : env -> evar_map -> constr -> (constr,constr) kind_of_term +val find_conclusion : env -> evar_map -> constr -> (constr, constr, Sorts.t, Univ.Instance.t) kind_of_term val is_arity : env -> evar_map -> constr -> bool val is_sort : env -> evar_map -> types -> bool -- cgit v1.2.3 From 3df2431a80f9817ce051334cb9c3b1f465bffb60 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 31 Mar 2017 23:20:25 +0200 Subject: Actually exporting delayed universes in the EConstr implementation. For now we only normalize sorts, and we leave instances for the next commit. --- engine/eConstr.ml | 54 +++++++++++++++++++++++++++++++++++--------- engine/eConstr.mli | 26 ++++++++++++++++++--- engine/namegen.ml | 2 +- engine/termops.ml | 9 ++++++-- plugins/ltac/rewrite.ml | 2 +- plugins/ltac/taccoerce.ml | 2 +- plugins/omega/coq_omega.ml | 2 +- pretyping/coercion.ml | 6 ++--- pretyping/constr_matching.ml | 13 +++++++---- pretyping/detyping.ml | 2 +- pretyping/evarconv.ml | 3 +++ pretyping/evardefine.ml | 6 ++--- pretyping/evarsolve.ml | 9 ++++++-- pretyping/inductiveops.ml | 1 + pretyping/pretyping.ml | 8 +++++-- pretyping/reductionops.mli | 6 ++--- pretyping/retyping.ml | 16 ++++++++----- pretyping/typing.ml | 15 +++++++----- pretyping/unification.ml | 2 ++ tactics/class_tactics.ml | 6 ++++- vernac/command.ml | 4 ++-- 21 files changed, 140 insertions(+), 54 deletions(-) diff --git a/engine/eConstr.ml b/engine/eConstr.ml index 094841d69..166340b41 100644 --- a/engine/eConstr.ml +++ b/engine/eConstr.ml @@ -16,12 +16,19 @@ 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 type t -val kind : Evd.evar_map -> t -> (t, t, Sorts.t, Univ.Instance.t) Constr.kind_of_term +val kind : Evd.evar_map -> t -> (t, t, ESorts.t, Univ.Instance.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, Sorts.t, Univ.Instance.t) Constr.kind_of_term -> t +val of_kind : (t, t, ESorts.t, Univ.Instance.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 @@ -33,6 +40,16 @@ val of_rel_decl : (Constr.t, Constr.types) Context.Rel.Declaration.pt -> (t, t) 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 + type t = Constr.t let safe_evar_value sigma ev = @@ -46,9 +63,6 @@ let rec whd_evar sigma c = | Some c -> whd_evar sigma c | None -> 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') @@ -82,8 +96,25 @@ let of_constr c = c let unsafe_to_constr c = c let unsafe_eq = Refl -let rec to_constr sigma t = - Constr.map (fun t -> to_constr sigma t) (whd_evar sigma t) +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 @@ -108,14 +139,14 @@ type rel_context = (constr, types) Context.Rel.pt let in_punivs a = (a, Univ.Instance.empty) -let mkProp = of_kind (Sort Sorts.prop) -let mkSet = of_kind (Sort Sorts.set) -let mkType u = of_kind (Sort (Sorts.Type u)) +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 s) +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)) @@ -730,6 +761,7 @@ let fresh_global ?loc ?rigid ?names env sigma reference = module Unsafe = struct +let to_sorts = ESorts.unsafe_to_sorts let to_constr = unsafe_to_constr let to_rel_decl = unsafe_to_rel_decl let to_named_decl = unsafe_to_named_decl diff --git a/engine/eConstr.mli b/engine/eConstr.mli index e7287fc06..db10fbf42 100644 --- a/engine/eConstr.mli +++ b/engine/eConstr.mli @@ -28,9 +28,26 @@ 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 + (** {5 Destructors} *) -val kind : Evd.evar_map -> t -> (t, t, Sorts.t, Univ.Instance.t) Constr.kind_of_term +val kind : Evd.evar_map -> t -> (t, t, ESorts.t, Univ.Instance.t) Constr.kind_of_term (** Same as {!Constr.kind} except that it expands evars and normalizes universes on the fly. *) @@ -43,7 +60,7 @@ val kind_of_type : Evd.evar_map -> t -> (t, t) Term.kind_of_type (** {5 Constructors} *) -val of_kind : (t, t, Sorts.t, Univ.Instance.t) Constr.kind_of_term -> t +val of_kind : (t, t, ESorts.t, Univ.Instance.t) Constr.kind_of_term -> t (** Construct a term from a view. *) val of_constr : Constr.t -> t @@ -123,7 +140,7 @@ 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 -> Sorts.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 @@ -242,6 +259,9 @@ sig 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 eq : (t, Constr.t) eq (** Use for transparent cast between types. *) end diff --git a/engine/namegen.ml b/engine/namegen.ml index 7b7302957..3b979f206 100644 --- a/engine/namegen.ml +++ b/engine/namegen.ml @@ -123,7 +123,7 @@ let hdchar env sigma 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 diff --git a/engine/termops.ml b/engine/termops.ml index 410fb75c5..a67916345 100644 --- a/engine/termops.ml +++ b/engine/termops.ml @@ -1181,14 +1181,19 @@ let base_sort_cmp pb s0 s1 = | _ -> false let rec is_Prop sigma c = match EConstr.kind sigma c with - | Sort (Prop Null) -> true + | 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 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 s1 s2 + 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 | _ -> EConstr.compare_constr sigma (fun t1 t2 -> f Reduction.CONV t1 t2) t1 t2 diff --git a/plugins/ltac/rewrite.ml b/plugins/ltac/rewrite.ml index 55172cba6..b84be4600 100644 --- a/plugins/ltac/rewrite.ml +++ b/plugins/ltac/rewrite.ml @@ -434,7 +434,7 @@ module TypeGlobal = struct end let sort_of_rel env evm rel = - Reductionops.sort_of_arity env evm (Retyping.get_type_of env evm rel) + ESorts.kind evm (Reductionops.sort_of_arity env evm (Retyping.get_type_of env evm rel)) let is_applied_rewrite_relation = PropGlobal.is_applied_rewrite_relation diff --git a/plugins/ltac/taccoerce.ml b/plugins/ltac/taccoerce.ml index 114b8dda0..95620b374 100644 --- a/plugins/ltac/taccoerce.ml +++ b/plugins/ltac/taccoerce.ml @@ -164,7 +164,7 @@ let id_of_name = function basename | Sort s -> begin - match s with + match ESorts.kind sigma s with | Prop _ -> Label.to_id (Label.make "Prop") | Type _ -> Label.to_id (Label.make "Type") end diff --git a/plugins/omega/coq_omega.ml b/plugins/omega/coq_omega.ml index adf926958..a2016cb03 100644 --- a/plugins/omega/coq_omega.ml +++ b/plugins/omega/coq_omega.ml @@ -1714,7 +1714,7 @@ let onClearedName2 id tac = end }) let rec is_Prop sigma c = match EConstr.kind sigma c with - | Sort (Prop Null) -> true + | Sort s -> Sorts.is_prop (ESorts.kind sigma s) | Cast (c,_,_) -> is_Prop sigma c | _ -> false diff --git a/pretyping/coercion.ml b/pretyping/coercion.ml index 8794f238b..542db7fdf 100644 --- a/pretyping/coercion.ml +++ b/pretyping/coercion.ml @@ -199,7 +199,7 @@ and coerce loc env evdref (x : EConstr.constr) (y : EConstr.constr) in match (EConstr.kind !evdref x, EConstr.kind !evdref y) with | Sort s, Sort s' -> - (match s, s' with + (match ESorts.kind !evdref s, ESorts.kind !evdref s' with | Prop x, Prop y when x == y -> None | Prop _, Type _ -> None | Type x, Type y when Univ.Universe.equal x y -> None (* false *) @@ -406,7 +406,7 @@ let inh_app_fun resolve_tc env evd j = let type_judgment env sigma j = match EConstr.kind sigma (whd_all env sigma j.uj_type) with - | Sort s -> {utj_val = j.uj_val; utj_type = s } + | Sort s -> {utj_val = j.uj_val; utj_type = ESorts.kind sigma s } | _ -> error_not_a_type env sigma j let inh_tosort_force loc env evd j = @@ -421,7 +421,7 @@ let inh_tosort_force loc env evd j = let inh_coerce_to_sort loc env evd j = let typ = whd_all env evd j.uj_type in match EConstr.kind evd typ with - | Sort s -> (evd,{ utj_val = j.uj_val; utj_type = s }) + | Sort s -> (evd,{ utj_val = j.uj_val; utj_type = ESorts.kind evd s }) | Evar ev -> let (evd',s) = Evardefine.define_evar_as_sort env evd ev in (evd',{ utj_val = j.uj_val; utj_type = s }) diff --git a/pretyping/constr_matching.ml b/pretyping/constr_matching.ml index cad21543b..30b83cf88 100644 --- a/pretyping/constr_matching.ml +++ b/pretyping/constr_matching.ml @@ -205,11 +205,14 @@ let matches_core env sigma convert allow_partial_app allow_bound_rels | PRel n1, Rel n2 when Int.equal n1 n2 -> subst - | PSort GProp, Sort (Prop Null) -> subst - - | PSort GSet, Sort (Prop Pos) -> subst - - | PSort (GType _), Sort (Type _) -> subst + | PSort ps, Sort s -> + + begin match ps, ESorts.kind sigma s with + | GProp, Prop Null -> subst + | GSet, Prop Pos -> subst + | GType _, Type _ -> subst + | _ -> raise PatternMatchingFailure + end | PApp (p, [||]), _ -> sorec ctx env subst p t diff --git a/pretyping/detyping.ml b/pretyping/detyping.ml index 1adda14ab..84022f57f 100644 --- a/pretyping/detyping.ml +++ b/pretyping/detyping.ml @@ -457,7 +457,7 @@ let rec detype flags avoid env sigma t = | Var id -> (try let _ = Global.lookup_named id in GRef (dl, VarRef id, None) with Not_found -> GVar (dl, id)) - | Sort s -> GSort (dl,detype_sort sigma s) + | Sort s -> GSort (dl,detype_sort sigma (ESorts.kind sigma s)) | Cast (c1,REVERTcast,c2) when not !Flags.raw_print -> detype flags avoid env sigma c1 | Cast (c1,k,c2) -> diff --git a/pretyping/evarconv.ml b/pretyping/evarconv.ml index 85cc8762e..b6fa25769 100644 --- a/pretyping/evarconv.ml +++ b/pretyping/evarconv.ml @@ -150,6 +150,7 @@ let check_conv_record env sigma (t1,sk1) (t2,sk2) = (Stack.append_app [|a;pop b|] Stack.empty) else raise Not_found | Sort s -> + let s = ESorts.kind sigma s in lookup_canonical_conversion (proji, Sort_cs (family_of_sort s)),[] | _ -> @@ -775,6 +776,8 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty | Sort s1, Sort s2 when app_empty -> (try + let s1 = ESorts.kind evd s1 in + let s2 = ESorts.kind evd s2 in let evd' = if pbty == CONV then Evd.set_eq_sort env evd s1 s2 diff --git a/pretyping/evardefine.ml b/pretyping/evardefine.ml index 20d86f81b..c5ae684e3 100644 --- a/pretyping/evardefine.ml +++ b/pretyping/evardefine.ml @@ -91,7 +91,7 @@ let define_pure_evar_as_product evd evk = let newenv = push_named (LocalAssum (id, dom)) evenv in let src = evar_source evk evd1 in let filter = Filter.extend 1 (evar_filter evi) in - if is_prop_sort s then + if is_prop_sort (ESorts.kind evd1 s) then (* Impredicative product, conclusion must fall in [Prop]. *) new_evar_unsafe newenv evd1 concl ~src ~filter else @@ -102,7 +102,7 @@ let define_pure_evar_as_product evd evk = (Sigma.to_evar_map evd3, e) in let prods = Univ.sup (univ_of_sort u1) (univ_of_sort srng) in - let evd3 = Evd.set_leq_sort evenv evd3 (Type prods) s in + let evd3 = Evd.set_leq_sort evenv evd3 (Type prods) (ESorts.kind evd1 s) in evd3, rng in let prod = mkProd (Name id, dom, subst_var id rng) in @@ -174,7 +174,7 @@ let define_evar_as_sort env evd (ev,args) = let concl = Reductionops.whd_all (evar_env evi) evd (EConstr.of_constr evi.evar_concl) in let sort = destSort evd concl in let evd' = Evd.define ev (Constr.mkSort s) evd in - Evd.set_leq_sort env evd' (Type (Univ.super u)) sort, s + Evd.set_leq_sort env evd' (Type (Univ.super u)) (ESorts.kind evd' sort), s (* Propagation of constraints through application and abstraction: Given a type constraint on a functional term, returns the type diff --git a/pretyping/evarsolve.ml b/pretyping/evarsolve.ml index 4d78d2eb0..77086d046 100644 --- a/pretyping/evarsolve.ml +++ b/pretyping/evarsolve.ml @@ -50,6 +50,7 @@ let refresh_universes ?(status=univ_rigid) ?(onlyalg=false) ?(refreshset=false) let modified = ref false in (* direction: true for fresh universes lower than the existing ones *) let refresh_sort status ~direction s = + let s = ESorts.kind !evdref s in let s' = evd_comb0 (new_sort_variable status) evdref in let evd = if direction then set_leq_sort env !evdref s' s @@ -59,7 +60,9 @@ let refresh_universes ?(status=univ_rigid) ?(onlyalg=false) ?(refreshset=false) in let rec refresh ~onlyalg status ~direction t = match EConstr.kind !evdref t with - | Sort (Type u as s) -> + | Sort s -> + begin match ESorts.kind !evdref s with + | Type u -> (match Univ.universe_level u with | None -> refresh_sort status ~direction s | Some l -> @@ -71,10 +74,12 @@ let refresh_universes ?(status=univ_rigid) ?(onlyalg=false) ?(refreshset=false) if onlyalg && alg then (evdref := Evd.make_flexible_variable !evdref false l; t) else t)) - | Sort (Prop Pos as s) when refreshset && not direction -> + | Prop Pos when refreshset && not direction -> (* Cannot make a universe "lower" than "Set", only refreshing when we want higher universes. *) refresh_sort status ~direction s + | _ -> t + end | Prod (na,u,v) -> mkProd (na, u, refresh ~onlyalg status ~direction v) | _ -> t diff --git a/pretyping/inductiveops.ml b/pretyping/inductiveops.ml index d5967c4bf..88c492f03 100644 --- a/pretyping/inductiveops.ml +++ b/pretyping/inductiveops.ml @@ -617,6 +617,7 @@ let type_of_inductive_knowing_conclusion env sigma ((mib,mip),u) conclty = | RegularArity s -> sigma, EConstr.of_constr (subst_instance_constr u s.mind_user_arity) | TemplateArity ar -> let _,scl = splay_arity env sigma conclty in + let scl = EConstr.ESorts.kind sigma scl in let ctx = List.rev mip.mind_arity_ctxt in let evdref = ref sigma in let ctx = diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml index 846d8055a..c673851c8 100644 --- a/pretyping/pretyping.ml +++ b/pretyping/pretyping.ml @@ -1071,7 +1071,11 @@ and pretype_instance k0 resolve_tc env evdref lvar loc hyps evk update = and pretype_type k0 resolve_tc valcon (env : ExtraEnv.t) evdref lvar = function | GHole (loc, knd, naming, None) -> let rec is_Type c = match EConstr.kind !evdref c with - | Sort (Type _) -> true + | Sort s -> + begin match ESorts.kind !evdref s with + | Type _ -> true + | Prop _ -> false + end | Cast (c, _, _) -> is_Type c | _ -> false in @@ -1081,7 +1085,7 @@ and pretype_type k0 resolve_tc valcon (env : ExtraEnv.t) evdref lvar = function let sigma = !evdref in let t = Retyping.get_type_of env.ExtraEnv.env sigma v in match EConstr.kind sigma (whd_all env.ExtraEnv.env sigma t) with - | Sort s -> s + | Sort s -> ESorts.kind sigma s | Evar ev when is_Type (existential_type sigma ev) -> evd_comb1 (define_evar_as_sort env.ExtraEnv.env) evdref ev | _ -> anomaly (Pp.str "Found a type constraint which is not a type") diff --git a/pretyping/reductionops.mli b/pretyping/reductionops.mli index 18416b142..01707b47a 100644 --- a/pretyping/reductionops.mli +++ b/pretyping/reductionops.mli @@ -215,8 +215,8 @@ val hnf_lam_applist : env -> evar_map -> constr -> constr list -> constr val splay_prod : env -> evar_map -> constr -> (Name.t * constr) list * constr val splay_lam : env -> evar_map -> constr -> (Name.t * constr) list * constr -val splay_arity : env -> evar_map -> constr -> (Name.t * constr) list * sorts -val sort_of_arity : env -> evar_map -> constr -> sorts +val splay_arity : env -> evar_map -> constr -> (Name.t * constr) list * ESorts.t +val sort_of_arity : env -> evar_map -> constr -> ESorts.t val splay_prod_n : env -> evar_map -> int -> constr -> rel_context * constr val splay_lam_n : env -> evar_map -> int -> constr -> rel_context * constr val splay_prod_assum : @@ -232,7 +232,7 @@ type 'a miota_args = { val reducible_mind_case : evar_map -> constr -> bool val reduce_mind_case : evar_map -> constr miota_args -> constr -val find_conclusion : env -> evar_map -> constr -> (constr, constr, Sorts.t, Univ.Instance.t) kind_of_term +val find_conclusion : env -> evar_map -> constr -> (constr, constr, ESorts.t, Univ.Instance.t) kind_of_term val is_arity : env -> evar_map -> constr -> bool val is_sort : env -> evar_map -> types -> bool diff --git a/pretyping/retyping.ml b/pretyping/retyping.ml index bb1b2901e..9c9751af8 100644 --- a/pretyping/retyping.ml +++ b/pretyping/retyping.ml @@ -77,7 +77,7 @@ let sort_of_atomic_type env sigma ft args = let rec concl_of_arity env n ar args = match EConstr.kind sigma (whd_all env sigma ar), args with | Prod (na, t, b), h::l -> concl_of_arity (push_rel (LocalDef (na, lift n h, t)) env) (n + 1) b l - | Sort s, [] -> s + | Sort s, [] -> ESorts.kind sigma s | _ -> retype_error NotASort in concl_of_arity env 0 ft (Array.to_list args) @@ -87,9 +87,11 @@ let type_of_var env id = let decomp_sort env sigma t = match EConstr.kind sigma (whd_all env sigma t) with - | Sort s -> s + | Sort s -> ESorts.kind sigma s | _ -> retype_error NotASort +let destSort sigma s = ESorts.kind sigma (destSort sigma s) + let retype ?(polyprop=true) sigma = let rec type_of env cstr = match EConstr.kind sigma cstr with @@ -142,8 +144,11 @@ let retype ?(polyprop=true) sigma = and sort_of env t = match EConstr.kind sigma t with | Cast (c,_, s) when isSort sigma s -> destSort sigma s - | Sort (Prop c) -> type1_sort - | Sort (Type u) -> Type (Univ.super u) + | Sort s -> + begin match ESorts.kind sigma s with + | Prop _ -> type1_sort + | Type u -> Type (Univ.super u) + end | Prod (name,t,c2) -> (match (sort_of env t, sort_of (push_rel (LocalAssum (name,t)) env) c2) with | _, (Prop Null as s) -> s @@ -163,8 +168,7 @@ let retype ?(polyprop=true) sigma = and sort_family_of env t = match EConstr.kind sigma t with | Cast (c,_, s) when isSort sigma s -> family_of_sort (destSort sigma s) - | Sort (Prop c) -> InType - | Sort (Type u) -> InType + | Sort _ -> InType | Prod (name,t,c2) -> let s2 = sort_family_of (push_rel (LocalAssum (name,t)) env) c2 in if not (is_impredicative_set env) && diff --git a/pretyping/typing.ml b/pretyping/typing.ml index dec22ecd0..d9d64e7eb 100644 --- a/pretyping/typing.ml +++ b/pretyping/typing.ml @@ -46,7 +46,7 @@ let inductive_type_knowing_parameters env sigma (ind,u) jl = let e_type_judgment env evdref j = match EConstr.kind !evdref (whd_all env !evdref j.uj_type) with - | Sort s -> {utj_val = j.uj_val; utj_type = s } + | Sort s -> {utj_val = j.uj_val; utj_type = ESorts.kind !evdref s } | Evar ev -> let (evd,s) = Evardefine.define_evar_as_sort env !evdref ev in evdref := evd; { utj_val = j.uj_val; utj_type = s } @@ -102,6 +102,7 @@ let e_is_correct_arity env evdref c pj ind specif params = if not (Evarconv.e_cumul env evdref a1 a1') then error (); srec (push_rel (LocalAssum (na1,a1)) env) t ar' | Sort s, [] -> + let s = ESorts.kind !evdref s in if not (Sorts.List.mem (Sorts.family s) allowed_sorts) then error () | Evar (ev,_), [] -> @@ -161,7 +162,7 @@ let check_type_fixpoint loc env evdref lna lar vdefj = (* FIXME: might depend on the level of actual parameters!*) let check_allowed_sort env sigma ind c p = let pj = Retyping.get_judgment_of env sigma p in - let ksort = family_of_sort (sort_of_arity env sigma pj.uj_type) in + let ksort = family_of_sort (ESorts.kind sigma (sort_of_arity env sigma pj.uj_type)) in let specif = Global.lookup_inductive (fst ind) in let sorts = elim_sorts specif in if not (List.exists ((==) ksort) sorts) then @@ -288,11 +289,13 @@ let rec execute env evdref cstr = check_cofix env !evdref cofix; make_judge (mkCoFix cofix) tys.(i) - | Sort (Prop c) -> - judge_of_prop_contents c - - | Sort (Type u) -> + | Sort s -> + begin match ESorts.kind !evdref s with + | Prop c -> + judge_of_prop_contents c + | Type u -> judge_of_type u + end | Proj (p, c) -> let cj = execute env evdref c in diff --git a/pretyping/unification.ml b/pretyping/unification.ml index 611d165fe..035b0c223 100644 --- a/pretyping/unification.ml +++ b/pretyping/unification.ml @@ -780,6 +780,8 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst : subst0) conv_at_top e else error_cannot_unify_local curenv sigma (m,n,cN) | Sort s1, Sort s2 -> (try + let s1 = ESorts.kind sigma s1 in + let s2 = ESorts.kind sigma s2 in let sigma' = if pb == CUMUL then Evd.set_leq_sort curenv sigma s1 s2 diff --git a/tactics/class_tactics.ml b/tactics/class_tactics.ml index c53e47d92..7eadde78d 100644 --- a/tactics/class_tactics.ml +++ b/tactics/class_tactics.ml @@ -512,7 +512,11 @@ let pr_depth l = prlist_with_sep (next_sep debug_seps) int (List.rev l) let is_Prop env sigma concl = let ty = Retyping.get_type_of env sigma concl in match EConstr.kind sigma ty with - | Sort (Prop Null) -> true + | Sort s -> + begin match ESorts.kind sigma s with + | Prop Null -> true + | _ -> false + end | _ -> false let is_unique env sigma concl = diff --git a/vernac/command.ml b/vernac/command.ml index 32b0abb8a..db203b425 100644 --- a/vernac/command.ml +++ b/vernac/command.ml @@ -987,8 +987,8 @@ let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation = try let ctx, ar = Reductionops.splay_prod_n env !evdref 2 relty in match ctx, EConstr.kind !evdref ar with - | [LocalAssum (_,t); LocalAssum (_,u)], Sort (Prop Null) - when Reductionops.is_conv env !evdref t u -> t + | [LocalAssum (_,t); LocalAssum (_,u)], Sort s + when Sorts.is_prop (ESorts.kind !evdref s) && Reductionops.is_conv env !evdref t u -> t | _, _ -> error () with e when CErrors.noncritical e -> error () in -- cgit v1.2.3 From 7babf0d42af11f5830bc157a671bd81b478a4f02 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sat, 1 Apr 2017 02:36:16 +0200 Subject: Using delayed universe instances in EConstr. The transition has been done a bit brutally. I think we can still save a lot of useless normalizations here and there by providing the right API in EConstr. Nonetheless, this is a first step. --- engine/eConstr.ml | 35 ++++++++++++++++--------- engine/eConstr.mli | 33 +++++++++++++++++------- engine/termops.ml | 14 ++++++---- engine/termops.mli | 4 +-- plugins/cc/cctac.ml | 6 ++++- plugins/firstorder/formula.ml | 6 ++++- plugins/funind/functional_principles_types.ml | 3 ++- plugins/funind/glob_term_to_relation.ml | 5 ++-- plugins/funind/indfun.ml | 11 +++++--- plugins/funind/invfun.ml | 5 ++-- plugins/funind/merge.ml | 7 ++--- plugins/funind/recdef.ml | 5 ++-- plugins/ltac/extratactics.ml4 | 2 +- plugins/ltac/taccoerce.ml | 2 +- plugins/quote/quote.ml | 5 ++-- pretyping/cases.ml | 5 ++-- pretyping/classops.ml | 9 ++++--- pretyping/classops.mli | 2 +- pretyping/detyping.ml | 2 ++ pretyping/evarconv.ml | 4 +-- pretyping/inductiveops.ml | 4 ++- pretyping/inductiveops.mli | 10 ++++---- pretyping/reductionops.ml | 17 +++++++----- pretyping/reductionops.mli | 2 +- pretyping/retyping.ml | 30 ++++++++++++---------- pretyping/tacred.ml | 37 +++++++++++++++------------ pretyping/tacred.mli | 6 ++--- pretyping/typeclasses.ml | 1 + pretyping/typeclasses.mli | 4 +-- pretyping/typing.ml | 33 ++++++++++++++---------- pretyping/unification.ml | 1 + proofs/logic.ml | 3 ++- proofs/tacmach.mli | 6 ++--- tactics/eauto.ml | 5 ++-- tactics/equality.ml | 9 ++++--- tactics/hipattern.ml | 17 ++++++------ tactics/hipattern.mli | 8 +++--- tactics/tacticals.ml | 11 +++++--- tactics/tacticals.mli | 6 ++--- tactics/tactics.ml | 22 +++++++++------- vernac/auto_ind_decl.ml | 20 ++++++++------- vernac/classes.ml | 5 ++-- 42 files changed, 257 insertions(+), 165 deletions(-) diff --git a/engine/eConstr.ml b/engine/eConstr.ml index 166340b41..28e9ffdb2 100644 --- a/engine/eConstr.ml +++ b/engine/eConstr.ml @@ -23,12 +23,21 @@ 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, Univ.Instance.t) Constr.kind_of_term +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, Univ.Instance.t) Constr.kind_of_term -> 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 @@ -50,6 +59,16 @@ struct let unsafe_to_sorts s = s end +module EInstance = +struct + type t = Univ.Instance.t + let make i = i + let kind sigma i = 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 = @@ -63,15 +82,6 @@ let rec whd_evar sigma c = | Some c -> whd_evar sigma c | None -> c end - | 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') | App (f, args) when isEvar f -> (** Enforce smart constructor invariant on applications *) let ev = destEvar f in @@ -137,7 +147,7 @@ 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, Univ.Instance.empty) +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)) @@ -762,6 +772,7 @@ let fresh_global ?loc ?rigid ?names env sigma reference = 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 diff --git a/engine/eConstr.mli b/engine/eConstr.mli index db10fbf42..3a9469e55 100644 --- a/engine/eConstr.mli +++ b/engine/eConstr.mli @@ -45,9 +45,21 @@ sig 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, Univ.Instance.t) Constr.kind_of_term +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. *) @@ -60,7 +72,7 @@ val kind_of_type : Evd.evar_map -> t -> (t, t) Term.kind_of_type (** {5 Constructors} *) -val of_kind : (t, t, ESorts.t, Univ.Instance.t) Constr.kind_of_term -> t +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 @@ -89,13 +101,13 @@ 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 : pconstant -> t +val mkConstU : constant * EInstance.t -> t val mkProj : (projection * t) -> t val mkInd : inductive -> t -val mkIndU : pinductive -> t +val mkIndU : inductive * EInstance.t -> t val mkConstruct : constructor -> t -val mkConstructU : pconstructor -> t -val mkConstructUi : pinductive * int -> 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 @@ -146,10 +158,10 @@ 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 puniverses +val destConst : Evd.evar_map -> t -> constant * EInstance.t val destEvar : Evd.evar_map -> t -> t pexistential -val destInd : Evd.evar_map -> t -> inductive puniverses -val destConstruct : Evd.evar_map -> t -> constructor puniverses +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 @@ -262,6 +274,9 @@ sig 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/termops.ml b/engine/termops.ml index a67916345..64f4c6dc5 100644 --- a/engine/termops.ml +++ b/engine/termops.ml @@ -1150,7 +1150,7 @@ let global_of_constr sigma c = | Const (c, u) -> ConstRef c, u | Ind (i, u) -> IndRef i, u | Construct (c, u) -> ConstructRef c, u - | Var id -> VarRef id, Univ.Instance.empty + | Var id -> VarRef id, EConstr.EInstance.empty | _ -> raise Not_found let is_global sigma c t = @@ -1169,8 +1169,12 @@ let isGlobalRef sigma c = let is_template_polymorphic env sigma f = match EConstr.kind sigma f with - | Ind ind -> Environ.template_polymorphic_pind ind env - | Const c -> Environ.template_polymorphic_pconstant c env + | 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 = @@ -1453,8 +1457,8 @@ let global_app_of_constr sigma c = | 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, Instance.empty), None - | Proj (p, c) -> (ConstRef (Projection.constant p), Instance.empty), Some c + | 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 = diff --git a/engine/termops.mli b/engine/termops.mli index 0dd5d96fb..fe6dfb0ce 100644 --- a/engine/termops.mli +++ b/engine/termops.mli @@ -250,7 +250,7 @@ val clear_named_body : Id.t -> env -> env 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 puniverses * constr option +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 *) @@ -259,7 +259,7 @@ val dependency_closure : env -> Evd.evar_map -> named_context -> Id.Set.t -> Id. (** Test if an identifier is the basename of a global reference *) val is_section_variable : Id.t -> bool -val global_of_constr : Evd.evar_map -> constr -> Globnames.global_reference puniverses +val global_of_constr : Evd.evar_map -> constr -> Globnames.global_reference * EInstance.t val is_global : Evd.evar_map -> Globnames.global_reference -> constr -> bool diff --git a/plugins/cc/cctac.ml b/plugins/cc/cctac.ml index c9c904e35..2d9dec095 100644 --- a/plugins/cc/cctac.ml +++ b/plugins/cc/cctac.ml @@ -66,6 +66,7 @@ let rec decompose_term env sigma t= decompose_term env sigma b) | Construct c -> let (((mind,i_ind),i_con),u)= c in + let u = EInstance.kind sigma u in let canon_mind = mind_of_kn (canonical_mind mind) in let canon_ind = canon_mind,i_ind in let (oib,_)=Global.lookup_inductive (canon_ind) in @@ -75,9 +76,11 @@ let rec decompose_term env sigma t= ci_nhyps=nargs-oib.mind_nparams} | Ind c -> let (mind,i_ind),u = c in + let u = EInstance.kind sigma u in let canon_mind = mind_of_kn (canonical_mind mind) in let canon_ind = canon_mind,i_ind in (Symb (Constr.mkIndU (canon_ind,u))) | Const (c,u) -> + let u = EInstance.kind sigma u in let canon_const = constant_of_kn (canonical_con c) in (Symb (Constr.mkConstU (canon_const,u))) | Proj (p, c) -> @@ -408,7 +411,8 @@ let build_term_to_complete uf meta pac = let real_args = List.map (fun i -> constr_of_term (term uf i)) pac.args in let dummy_args = List.rev (List.init pac.arity meta) in let all_args = List.rev_append real_args dummy_args in - applist (mkConstructU cinfo.ci_constr, all_args) + let (kn, u) = cinfo.ci_constr in + applist (mkConstructU (kn, EInstance.make u), all_args) let cc_tactic depth additionnal_terms = Proofview.Goal.enter { enter = begin fun gl -> diff --git a/plugins/firstorder/formula.ml b/plugins/firstorder/formula.ml index 87bac2fe3..7773f6a2f 100644 --- a/plugins/firstorder/formula.ml +++ b/plugins/firstorder/formula.ml @@ -91,6 +91,7 @@ let kind_of_formula gl term = Some (i,l,n)-> let l = List.map EConstr.Unsafe.to_constr l in let ind,u=EConstr.destInd (project gl) i in + let u = EConstr.EInstance.kind (project gl) u in let (mib,mip) = Global.lookup_inductive ind in let nconstr=Array.length mip.mind_consnames in if Int.equal nconstr 0 then @@ -112,7 +113,10 @@ let kind_of_formula gl term = Or((ind,u),l,is_trivial) | _ -> match match_with_sigma_type (project gl) (EConstr.of_constr cciterm) with - Some (i,l)-> Exists((EConstr.destInd (project gl) i),List.map EConstr.Unsafe.to_constr l) + Some (i,l)-> + let (ind, u) = EConstr.destInd (project gl) i in + let u = EConstr.EInstance.kind (project gl) u in + Exists((ind, u), List.map EConstr.Unsafe.to_constr l) |_-> Atom (normalize cciterm) type atoms = {positive:constr list;negative:constr list} diff --git a/plugins/funind/functional_principles_types.ml b/plugins/funind/functional_principles_types.ml index e845db3bc..529b91c4c 100644 --- a/plugins/funind/functional_principles_types.ml +++ b/plugins/funind/functional_principles_types.ml @@ -300,7 +300,8 @@ let build_functional_principle (evd:Evd.evar_map ref) interactive_proof old_prin hook ; (* let _tim1 = System.get_time () in *) - ignore (Pfedit.by (Proofview.V82.tactic (proof_tac (Array.map EConstr.mkConstU funs) mutr_nparams))); + let map (c, u) = EConstr.mkConstU (c, EConstr.EInstance.make u) in + ignore (Pfedit.by (Proofview.V82.tactic (proof_tac (Array.map map funs) mutr_nparams))); (* let _tim2 = System.get_time () in *) (* begin *) (* let dur1 = System.time_difference tim1 tim2 in *) diff --git a/plugins/funind/glob_term_to_relation.ml b/plugins/funind/glob_term_to_relation.ml index 63bd3848f..8aab3b742 100644 --- a/plugins/funind/glob_term_to_relation.ml +++ b/plugins/funind/glob_term_to_relation.ml @@ -1275,8 +1275,9 @@ let do_build_inductive let open Context.Named.Declaration in let evd,env = Array.fold_right2 - (fun id c (evd,env) -> - let evd,t = Typing.type_of env evd (EConstr.mkConstU c) in + (fun id (c, u) (evd,env) -> + let u = EConstr.EInstance.make u in + let evd,t = Typing.type_of env evd (EConstr.mkConstU (c, u)) in let t = EConstr.Unsafe.to_constr t in evd, Environ.push_named (LocalAssum (id,t)) diff --git a/plugins/funind/indfun.ml b/plugins/funind/indfun.ml index a7489fb7b..2852152e1 100644 --- a/plugins/funind/indfun.ml +++ b/plugins/funind/indfun.ml @@ -249,7 +249,8 @@ let derive_inversion fix_names = Evd.fresh_global (Global.env ()) evd (Constrintern.locate_reference (Libnames.qualid_of_ident id)) in let c = EConstr.of_constr c in - evd, destConst evd c::l + let (cst, u) = destConst evd c in + evd, (cst, EInstance.kind evd u) :: l ) fix_names (evd',[]) @@ -412,7 +413,9 @@ let register_struct is_rec (fixpoint_exprl:(Vernacexpr.fixpoint_expr * Vernacexp Evd.fresh_global (Global.env ()) evd (Constrintern.locate_reference (Libnames.qualid_of_ident fname)) in let c = EConstr.of_constr c in - evd,((destConst evd c)::l) + let (cst, u) = destConst evd c in + let u = EInstance.kind evd u in + evd,((cst, u) :: l) ) (Evd.from_env (Global.env ()),[]) fixpoint_exprl @@ -427,7 +430,9 @@ let register_struct is_rec (fixpoint_exprl:(Vernacexpr.fixpoint_expr * Vernacexp Evd.fresh_global (Global.env ()) evd (Constrintern.locate_reference (Libnames.qualid_of_ident fname)) in let c = EConstr.of_constr c in - evd,((destConst evd c)::l) + let (cst, u) = destConst evd c in + let u = EInstance.kind evd u in + evd,((cst, u) :: l) ) (Evd.from_env (Global.env ()),[]) fixpoint_exprl diff --git a/plugins/funind/invfun.ml b/plugins/funind/invfun.ml index 2c2cd919b..94ec0a898 100644 --- a/plugins/funind/invfun.ml +++ b/plugins/funind/invfun.ml @@ -783,7 +783,8 @@ let derive_correctness make_scheme functional_induction (funs: pconstant list) ( assert (funs <> []); assert (graphs <> []); let funs = Array.of_list funs and graphs = Array.of_list graphs in - let funs_constr = Array.map mkConstU funs in + let map (c, u) = mkConstU (c, EInstance.make u) in + let funs_constr = Array.map map funs in States.with_state_protection_on_exception (fun () -> let env = Global.env () in @@ -882,7 +883,7 @@ let derive_correctness make_scheme functional_induction (funs: pconstant list) ( (Indrec.build_mutual_induction_scheme (Global.env ()) !evd (Array.to_list (Array.mapi - (fun i _ -> ((kn,i),u(* Univ.Instance.empty *)),true,InType) + (fun i _ -> ((kn,i), EInstance.kind !evd u),true,InType) mib.Declarations.mind_packets ) ) diff --git a/plugins/funind/merge.ml b/plugins/funind/merge.ml index 44aacaf45..c0298d06c 100644 --- a/plugins/funind/merge.ml +++ b/plugins/funind/merge.ml @@ -136,13 +136,14 @@ let prNamedRLDecl s lc = let showind (id:Id.t) = let cstrid = Constrintern.global_reference id in - let ind1,cstrlist = Inductiveops.find_inductive (Global.env()) Evd.empty (EConstr.of_constr cstrid) in - let mib1,ib1 = Inductive.lookup_mind_specif (Global.env()) (fst ind1) in + let (ind1, u),cstrlist = Inductiveops.find_inductive (Global.env()) Evd.empty (EConstr.of_constr cstrid) in + let mib1,ib1 = Inductive.lookup_mind_specif (Global.env()) ind1 in + let u = EConstr.Unsafe.to_instance u in List.iter (fun decl -> print_string (string_of_name (Context.Rel.Declaration.get_name decl) ^ ":"); prconstr (RelDecl.get_type decl); print_string "\n") ib1.mind_arity_ctxt; - Printf.printf "arity :"; prconstr (Inductiveops.type_of_inductive (Global.env ()) ind1); + Printf.printf "arity :"; prconstr (Inductiveops.type_of_inductive (Global.env ()) (ind1, u)); Array.iteri (fun i x -> Printf.printf"type constr %d :" i ; prconstr x) ib1.mind_user_lc diff --git a/plugins/funind/recdef.ml b/plugins/funind/recdef.ml index 0a288c76e..5460d6fb7 100644 --- a/plugins/funind/recdef.ml +++ b/plugins/funind/recdef.ml @@ -86,9 +86,10 @@ let def_of_const t = let type_of_const sigma t = match (EConstr.kind sigma t) with - | Const sp -> + | Const (sp, u) -> + let u = EInstance.kind sigma u in (* FIXME discarding universe constraints *) - Typeops.type_of_constant_in (Global.env()) sp + Typeops.type_of_constant_in (Global.env()) (sp, u) |_ -> assert false let constr_of_global x = diff --git a/plugins/ltac/extratactics.ml4 b/plugins/ltac/extratactics.ml4 index 55108631b..b34afd51b 100644 --- a/plugins/ltac/extratactics.ml4 +++ b/plugins/ltac/extratactics.ml4 @@ -1073,7 +1073,7 @@ let decompose l c = Proofview.Goal.enter { enter = begin fun gl -> let sigma = Tacmach.New.project gl in let to_ind c = - if isInd sigma c then Univ.out_punivs (destInd sigma c) + if isInd sigma c then fst (destInd sigma c) else error "not an inductive type" in let l = List.map to_ind l in diff --git a/plugins/ltac/taccoerce.ml b/plugins/ltac/taccoerce.ml index 95620b374..b76009c99 100644 --- a/plugins/ltac/taccoerce.ml +++ b/plugins/ltac/taccoerce.ml @@ -255,7 +255,7 @@ let coerce_to_evaluable_ref env sigma v = | IndRef _ | ConstructRef _ -> fail () else match Value.to_constr v with - | Some c when isConst sigma c -> EvalConstRef (Univ.out_punivs (destConst sigma c)) + | Some c when isConst sigma c -> EvalConstRef (fst (destConst sigma c)) | Some c when isVar sigma c -> EvalVarRef (destVar sigma c) | _ -> fail () in if Tacred.is_evaluable env ev then ev else fail () diff --git a/plugins/quote/quote.ml b/plugins/quote/quote.ml index 23069a9ab..fc9d70ae7 100644 --- a/plugins/quote/quote.ml +++ b/plugins/quote/quote.ml @@ -225,8 +225,9 @@ let compute_rhs env sigma bodyi index_of_f = let compute_ivs f cs gl = let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in - let cst = try destConst sigma f with DestKO -> i_can't_do_that () in - let body = Environ.constant_value_in (Global.env()) cst in + let (cst, u) = try destConst sigma f with DestKO -> i_can't_do_that () in + let u = EInstance.kind sigma u in + let body = Environ.constant_value_in (Global.env()) (cst, u) in let body = EConstr.of_constr body in match decomp_term sigma body with | Fix(([| len |], 0), ([| name |], [| typ |], [| body2 |])) -> diff --git a/pretyping/cases.ml b/pretyping/cases.ml index 38c105666..c5cf74ccf 100644 --- a/pretyping/cases.ml +++ b/pretyping/cases.ml @@ -279,6 +279,7 @@ let rec find_row_ind = function let inductive_template evdref env tmloc ind = let indu = evd_comb1 (Evd.fresh_inductive_instance env) evdref ind in let arsign = inductive_alldecls_env env indu in + let indu = on_snd EInstance.make indu in let hole_source i = match tmloc with | Some loc -> (loc, Evar_kinds.TomatchTypeParameter (ind,i)) | None -> (Loc.ghost, Evar_kinds.TomatchTypeParameter (ind,i)) in @@ -1314,7 +1315,7 @@ let build_branch initial current realargs deps (realnames,curname) pb arsign eqn let cur_alias = lift const_info.cs_nargs current in let ind = mkApp ( - applist (mkIndU (inductive_of_constructor (fst const_info.cs_cstr), snd const_info.cs_cstr), + applist (mkIndU (inductive_of_constructor (fst const_info.cs_cstr), EInstance.make (snd const_info.cs_cstr)), List.map (EConstr.of_constr %> lift const_info.cs_nargs) const_info.cs_params), Array.map EConstr.of_constr const_info.cs_concl_realargs) in Alias (initial,(aliasname,cur_alias,(ci,ind))) in @@ -2104,7 +2105,7 @@ let constr_of_pat env evdref arsign pat avoid = let args = List.rev args in let patargs = List.rev patargs in let pat' = PatCstr (l, cstr, patargs, alias) in - let cstr = mkConstructU ci.cs_cstr in + let cstr = mkConstructU (on_snd EInstance.make ci.cs_cstr) in let app = applist (cstr, List.map (lift (List.length sign)) params) in let app = applist (app, args) in let apptype = Retyping.get_type_of env ( !evdref) app in diff --git a/pretyping/classops.ml b/pretyping/classops.ml index 13310c44d..632ba0d9c 100644 --- a/pretyping/classops.ml +++ b/pretyping/classops.ml @@ -193,15 +193,16 @@ let coercion_exists coe = CoeTypMap.mem coe !coercion_tab (* find_class_type : evar_map -> constr -> cl_typ * universe_list * constr list *) let find_class_type sigma t = + let open EConstr in let t', args = Reductionops.whd_betaiotazeta_stack sigma t in match EConstr.kind sigma t' with - | Var id -> CL_SECVAR id, Univ.Instance.empty, args + | Var id -> CL_SECVAR id, EInstance.empty, args | Const (sp,u) -> CL_CONST sp, u, args | Proj (p, c) when not (Projection.unfolded p) -> - CL_PROJ (Projection.constant p), Univ.Instance.empty, (c :: args) + CL_PROJ (Projection.constant p), EInstance.empty, (c :: args) | Ind (ind_sp,u) -> CL_IND ind_sp, u, args - | Prod (_,_,_) -> CL_FUN, Univ.Instance.empty, [] - | Sort _ -> CL_SORT, Univ.Instance.empty, [] + | Prod (_,_,_) -> CL_FUN, EInstance.empty, [] + | Sort _ -> CL_SORT, EInstance.empty, [] | _ -> raise Not_found diff --git a/pretyping/classops.mli b/pretyping/classops.mli index a1d030f12..0d741a5a5 100644 --- a/pretyping/classops.mli +++ b/pretyping/classops.mli @@ -60,7 +60,7 @@ val class_info_from_index : cl_index -> cl_typ * cl_info_typ (** [find_class_type env sigma c] returns the head reference of [c], its universe instance and its arguments *) -val find_class_type : evar_map -> types -> cl_typ * Univ.universe_instance * constr list +val find_class_type : evar_map -> types -> cl_typ * EInstance.t * constr list (** raises [Not_found] if not convertible to a class *) val class_of : env -> evar_map -> types -> types * cl_index diff --git a/pretyping/detyping.ml b/pretyping/detyping.ml index 84022f57f..e4d7ab38d 100644 --- a/pretyping/detyping.ml +++ b/pretyping/detyping.ml @@ -438,6 +438,7 @@ let detype_level sigma l = GType (Some (dl, Pp.string_of_ppcmds (Termops.pr_evd_level sigma l))) let detype_instance sigma l = + let l = EInstance.kind sigma l in if Univ.Instance.is_empty l then None else Some (List.map (detype_level sigma) (Array.to_list (Univ.Instance.to_array l))) @@ -507,6 +508,7 @@ let rec detype flags avoid env sigma t = let ty = Retyping.get_type_of (snd env) sigma c in let ((ind,u), args) = Inductiveops.find_mrectype (snd env) sigma ty in let body' = strip_lam_assum body in + let u = EInstance.kind sigma u in let body' = CVars.subst_instance_constr u body' in let body' = EConstr.of_constr body' in substl (c :: List.rev args) body' diff --git a/pretyping/evarconv.ml b/pretyping/evarconv.ml index b6fa25769..9c9350ab1 100644 --- a/pretyping/evarconv.ml +++ b/pretyping/evarconv.ml @@ -51,9 +51,9 @@ let unfold_projection env evd ts p c = let eval_flexible_term ts env evd c = match EConstr.kind evd c with - | Const (c,u as cu) -> + | Const (c, u) -> if is_transparent_constant ts c - then Option.map EConstr.of_constr (constant_opt_value_in env cu) + then Option.map EConstr.of_constr (constant_opt_value_in env (c, EInstance.kind evd u)) else None | Rel n -> (try match lookup_rel n env with diff --git a/pretyping/inductiveops.ml b/pretyping/inductiveops.ml index 88c492f03..5b42add28 100644 --- a/pretyping/inductiveops.ml +++ b/pretyping/inductiveops.ml @@ -74,6 +74,7 @@ let substnl_ind_type l n = map_inductive_type (EConstr.Vars.substnl l n) let mkAppliedInd (IndType ((ind,params), realargs)) = let open EConstr in + let ind = on_snd EInstance.make ind in applist (mkIndU ind, (List.map EConstr.of_constr params)@realargs) (* Does not consider imbricated or mutually recursive types *) @@ -471,11 +472,12 @@ let find_rectype env sigma c = let open EConstr in let (t, l) = decompose_app sigma (whd_all env sigma c) in match EConstr.kind sigma t with - | Ind (ind,u as indu) -> + | Ind (ind,u) -> let (mib,mip) = Inductive.lookup_mind_specif env ind in if mib.mind_nparams > List.length l then raise Not_found; let l = List.map EConstr.Unsafe.to_constr l in let (par,rargs) = List.chop mib.mind_nparams l in + let indu = (ind, EInstance.kind sigma u) in IndType((indu, par),List.map EConstr.of_constr rargs) | _ -> raise Not_found diff --git a/pretyping/inductiveops.mli b/pretyping/inductiveops.mli index ab470a540..bdb6f996b 100644 --- a/pretyping/inductiveops.mli +++ b/pretyping/inductiveops.mli @@ -161,12 +161,12 @@ val make_arity : env -> evar_map -> bool -> inductive_family -> sorts -> EConstr val build_branch_type : env -> evar_map -> bool -> constr -> constructor_summary -> types (** Raise [Not_found] if not given a valid inductive type *) -val extract_mrectype : evar_map -> EConstr.t -> pinductive * EConstr.constr list -val find_mrectype : env -> evar_map -> EConstr.types -> pinductive * EConstr.constr list -val find_mrectype_vect : env -> evar_map -> EConstr.types -> pinductive * EConstr.constr array +val extract_mrectype : evar_map -> EConstr.t -> (inductive * EConstr.EInstance.t) * EConstr.constr list +val find_mrectype : env -> evar_map -> EConstr.types -> (inductive * EConstr.EInstance.t) * EConstr.constr list +val find_mrectype_vect : env -> evar_map -> EConstr.types -> (inductive * EConstr.EInstance.t) * EConstr.constr array val find_rectype : env -> evar_map -> EConstr.types -> inductive_type -val find_inductive : env -> evar_map -> EConstr.types -> pinductive * constr list -val find_coinductive : env -> evar_map -> EConstr.types -> pinductive * constr list +val find_inductive : env -> evar_map -> EConstr.types -> (inductive * EConstr.EInstance.t) * constr list +val find_coinductive : env -> evar_map -> EConstr.types -> (inductive * EConstr.EInstance.t) * constr list (********************) diff --git a/pretyping/reductionops.ml b/pretyping/reductionops.ml index 8be3b8328..270320538 100644 --- a/pretyping/reductionops.ml +++ b/pretyping/reductionops.ml @@ -550,7 +550,7 @@ struct let constr_of_cst_member f sk = match f with - | Cst_const (c, u) -> mkConstU (c,u), sk + | Cst_const (c, u) -> mkConstU (c, EInstance.make u), sk | Cst_proj p -> match decomp sk with | Some (hd, sk) -> mkProj (p, hd), sk @@ -703,7 +703,7 @@ let magicaly_constant_of_fixbody env sigma reference bd = function csts Univ.LMap.empty in let inst = Instance.subst_fn (fun u -> Univ.LMap.find u subst) u in - mkConstU (cst,inst) + mkConstU (cst, EInstance.make inst) | None -> bd end with @@ -856,7 +856,8 @@ let rec whd_state_gen ?csts ~refold ~tactic_mode flags env sigma = | Some body -> whrec cst_l (EConstr.of_constr body, stack) | None -> fold ()) | Const (c,u as const) when CClosure.RedFlags.red_set flags (CClosure.RedFlags.fCONST c) -> - (match constant_opt_value_in env const with + let u' = EInstance.kind sigma u in + (match constant_opt_value_in env (fst const, u') with | None -> fold () | Some body -> let body = EConstr.of_constr body in @@ -895,7 +896,7 @@ let rec whd_state_gen ?csts ~refold ~tactic_mode flags env sigma = | None -> fold () | Some (bef,arg,s') -> whrec Cst_stack.empty - (arg,Stack.Cst(Stack.Cst_const const,curr,remains,bef,cst_l)::s') + (arg,Stack.Cst(Stack.Cst_const (fst const, u'),curr,remains,bef,cst_l)::s') ) | Proj (p, c) when CClosure.RedFlags.red_projection flags p -> (let pb = lookup_projection p env in @@ -998,6 +999,7 @@ let rec whd_state_gen ?csts ~refold ~tactic_mode flags env sigma = (match constant_opt_value_in env const with | None -> fold () | Some body -> + let const = (fst const, EInstance.make (snd const)) in let body = EConstr.of_constr body in whrec (if refold then Cst_stack.add_cst (mkConstU const) cst_l else cst_l) (body, s' @ (Stack.append_app [|x'|] s''))) @@ -1657,12 +1659,13 @@ let meta_reducible_instance evd b = let head_unfold_under_prod ts env sigma c = - let unfold (cst,u as cstu) = + let unfold (cst,u) = + let cstu = (cst, EInstance.kind sigma u) in if Cpred.mem cst (snd ts) then match constant_opt_value_in env cstu with | Some c -> EConstr.of_constr c - | None -> mkConstU cstu - else mkConstU cstu in + | None -> mkConstU (cst, u) + else mkConstU (cst, u) in let rec aux c = match EConstr.kind sigma c with | Prod (n,t,c) -> mkProd (n,aux t, aux c) diff --git a/pretyping/reductionops.mli b/pretyping/reductionops.mli index 01707b47a..752c30a8a 100644 --- a/pretyping/reductionops.mli +++ b/pretyping/reductionops.mli @@ -232,7 +232,7 @@ type 'a miota_args = { val reducible_mind_case : evar_map -> constr -> bool val reduce_mind_case : evar_map -> constr miota_args -> constr -val find_conclusion : env -> evar_map -> constr -> (constr, constr, ESorts.t, Univ.Instance.t) kind_of_term +val find_conclusion : env -> evar_map -> constr -> (constr, constr, ESorts.t, EInstance.t) kind_of_term val is_arity : env -> evar_map -> constr -> bool val is_sort : env -> evar_map -> types -> bool diff --git a/pretyping/retyping.ml b/pretyping/retyping.ml index 9c9751af8..496c706ec 100644 --- a/pretyping/retyping.ml +++ b/pretyping/retyping.ml @@ -102,10 +102,10 @@ let retype ?(polyprop=true) sigma = let ty = RelDecl.get_type (lookup_rel n env) in lift n ty | Var id -> type_of_var env id - | Const cst -> EConstr.of_constr (rename_type_of_constant env cst) + | Const (cst, u) -> EConstr.of_constr (rename_type_of_constant env (cst, EInstance.kind sigma u)) | Evar ev -> existential_type sigma ev - | Ind ind -> EConstr.of_constr (rename_type_of_inductive env ind) - | Construct cstr -> EConstr.of_constr (rename_type_of_constructor env cstr) + | Ind (ind, u) -> EConstr.of_constr (rename_type_of_inductive env (ind, EInstance.kind sigma u)) + | Construct (cstr, u) -> EConstr.of_constr (rename_type_of_constructor env (cstr, EInstance.kind sigma u)) | Case (_,p,c,lf) -> let Inductiveops.IndType(indf,realargs) = let t = type_of env c in @@ -186,16 +186,20 @@ let retype ?(polyprop=true) sigma = let argtyps = Array.map (fun c -> lazy (EConstr.to_constr sigma (type_of env c))) args in match EConstr.kind sigma c with - | Ind ind -> - let mip = lookup_mind_specif env (fst ind) in + | Ind (ind, u) -> + let u = EInstance.kind sigma u in + let mip = lookup_mind_specif env ind in EConstr.of_constr (try Inductive.type_of_inductive_knowing_parameters - ~polyprop env (mip,snd ind) argtyps + ~polyprop env (mip, u) argtyps with Reduction.NotArity -> retype_error NotAnArity) - | Const cst -> - EConstr.of_constr (try Typeops.type_of_constant_knowing_parameters_in env cst argtyps + | Const (cst, u) -> + let u = EInstance.kind sigma u in + EConstr.of_constr (try Typeops.type_of_constant_knowing_parameters_in env (cst, u) argtyps with Reduction.NotArity -> retype_error NotAnArity) | Var id -> type_of_var env id - | Construct cstr -> EConstr.of_constr (type_of_constructor env cstr) + | Construct (cstr, u) -> + let u = EInstance.kind sigma u in + EConstr.of_constr (type_of_constructor env (cstr, u)) | _ -> assert false in type_of, sort_of, sort_family_of, @@ -212,13 +216,13 @@ let type_of_global_reference_knowing_conclusion env sigma c conclty = match EConstr.kind sigma c with | Ind (ind,u) -> let spec = Inductive.lookup_mind_specif env ind in - type_of_inductive_knowing_conclusion env sigma (spec,u) conclty - | Const cst -> - let t = constant_type_in env cst in + type_of_inductive_knowing_conclusion env sigma (spec, EInstance.kind sigma u) conclty + | Const (cst, u) -> + let t = constant_type_in env (cst, EInstance.kind sigma u) in (* TODO *) sigma, EConstr.of_constr (Typeops.type_of_constant_type_knowing_parameters env t [||]) | Var id -> sigma, type_of_var env id - | Construct cstr -> sigma, EConstr.of_constr (type_of_constructor env cstr) + | Construct (cstr, u) -> sigma, EConstr.of_constr (type_of_constructor env (cstr, EInstance.kind sigma u)) | _ -> assert false (* Profiling *) diff --git a/pretyping/tacred.ml b/pretyping/tacred.ml index ef9f39d77..67221046b 100644 --- a/pretyping/tacred.ml +++ b/pretyping/tacred.ml @@ -60,6 +60,7 @@ let is_evaluable env = function let value_of_evaluable_ref env evref u = match evref with | EvalConstRef con -> + let u = Unsafe.to_instance u in EConstr.of_constr (try constant_value_in env (con,u) with NotEvaluableConst IsProj -> raise (Invalid_argument "value_of_evaluable_ref")) @@ -103,9 +104,9 @@ let isEvalRef env sigma c = match EConstr.kind sigma c with let destEvalRefU sigma c = match EConstr.kind sigma c with | Const (cst,u) -> EvalConst cst, u - | Var id -> (EvalVar id, Univ.Instance.empty) - | Rel n -> (EvalRel n, Univ.Instance.empty) - | Evar ev -> (EvalEvar ev, Univ.Instance.empty) + | Var id -> (EvalVar id, EInstance.empty) + | Rel n -> (EvalRel n, EInstance.empty) + | Evar ev -> (EvalEvar ev, EInstance.empty) | _ -> anomaly (Pp.str "Not an unfoldable reference") let unsafe_reference_opt_value env sigma eval = @@ -125,7 +126,9 @@ let unsafe_reference_opt_value env sigma eval = let reference_opt_value env sigma eval u = match eval with - | EvalConst cst -> Option.map EConstr.of_constr (constant_opt_value_in env (cst,u)) + | EvalConst cst -> + let u = EInstance.kind sigma u in + Option.map EConstr.of_constr (constant_opt_value_in env (cst,u)) | EvalVar id -> env |> lookup_named id |> NamedDecl.get_value | EvalRel n -> @@ -519,13 +522,13 @@ let reduce_mind_case_use_function func env sigma mia = mutual inductive, try to reuse the global name if the block was indeed initially built as a global definition *) - let kn = map_puniverses (fun x -> con_with_label x (Label.of_id id)) - (destConst sigma func) - in - try match constant_opt_value_in env kn with + let (kn, u) = destConst sigma func in + let kn = con_with_label kn (Label.of_id id) in + let cst = (kn, EInstance.kind sigma u) in + try match constant_opt_value_in env cst with | None -> None (* TODO: check kn is correct *) - | Some _ -> Some (minargs,mkConstU kn) + | Some _ -> Some (minargs,mkConstU (kn, u)) with Not_found -> None else fun _ -> None in @@ -539,14 +542,15 @@ let match_eval_ref env sigma constr = match EConstr.kind sigma constr with | Const (sp, u) when is_evaluable env (EvalConstRef sp) -> Some (EvalConst sp, u) - | Var id when is_evaluable env (EvalVarRef id) -> Some (EvalVar id, Univ.Instance.empty) - | Rel i -> Some (EvalRel i, Univ.Instance.empty) - | Evar ev -> Some (EvalEvar ev, Univ.Instance.empty) + | Var id when is_evaluable env (EvalVarRef id) -> Some (EvalVar id, EInstance.empty) + | Rel i -> Some (EvalRel i, EInstance.empty) + | Evar ev -> Some (EvalEvar ev, EInstance.empty) | _ -> None let match_eval_ref_value env sigma constr = match EConstr.kind sigma constr with | Const (sp, u) when is_evaluable env (EvalConstRef sp) -> + let u = EInstance.kind sigma u in Some (EConstr.of_constr (constant_value_in env (sp, u))) | Var id when is_evaluable env (EvalVarRef id) -> env |> lookup_named id |> NamedDecl.get_value @@ -628,8 +632,9 @@ let whd_nothing_for_iota env sigma s = | Meta ev -> (try whrec (EConstr.of_constr (Evd.meta_value sigma ev), stack) with Not_found -> s) - | Const const when is_transparent_constant full_transparent_state (fst const) -> - (match constant_opt_value_in env const with + | Const (const, u) when is_transparent_constant full_transparent_state const -> + let u = EInstance.kind sigma u in + (match constant_opt_value_in env (const, u) with | Some body -> whrec (EConstr.of_constr body, stack) | None -> s) | LetIn (_,b,_,c) -> stacklam whrec [b] sigma c stack @@ -955,7 +960,7 @@ let simpl env sigma c = strong whd_simpl env sigma c let matches_head env sigma c t = match EConstr.kind sigma t with | App (f,_) -> Constr_matching.matches env sigma c f - | Proj (p, _) -> Constr_matching.matches env sigma c (mkConstU (Projection.constant p, Univ.Instance.empty)) + | Proj (p, _) -> Constr_matching.matches env sigma c (mkConstU (Projection.constant p, EInstance.empty)) | _ -> raise Constr_matching.PatternMatchingFailure (** FIXME: Specific function to handle projections: it ignores what happens on the @@ -1039,7 +1044,7 @@ let contextually byhead occs f env sigma t = let match_constr_evaluable_ref sigma c evref = match EConstr.kind sigma c, evref with | Const (c,u), EvalConstRef c' when eq_constant c c' -> Some u - | Var id, EvalVarRef id' when id_eq id id' -> Some Univ.Instance.empty + | Var id, EvalVarRef id' when id_eq id id' -> Some EInstance.empty | _, _ -> None let substlin env sigma evalref n (nowhere_except_in,locs) c = diff --git a/pretyping/tacred.mli b/pretyping/tacred.mli index a4499015d..76d0bc241 100644 --- a/pretyping/tacred.mli +++ b/pretyping/tacred.mli @@ -76,12 +76,12 @@ val cbv_norm_flags : CClosure.RedFlags.reds -> reduction_function (** [reduce_to_atomic_ind env sigma t] puts [t] in the form [t'=(I args)] with [I] an inductive definition; returns [I] and [t'] or fails with a user error *) -val reduce_to_atomic_ind : env -> evar_map -> types -> pinductive * types +val reduce_to_atomic_ind : env -> evar_map -> types -> (inductive * EInstance.t) * types (** [reduce_to_quantified_ind env sigma t] puts [t] in the form [t'=(x1:A1)..(xn:An)(I args)] with [I] an inductive definition; returns [I] and [t'] or fails with a user error *) -val reduce_to_quantified_ind : env -> evar_map -> types -> pinductive * types +val reduce_to_quantified_ind : env -> evar_map -> types -> (inductive * EInstance.t) * types (** [reduce_to_quantified_ref env sigma ref t] try to put [t] in the form [t'=(x1:A1)..(xn:An)(ref args)] and fails with user error if not possible *) @@ -92,7 +92,7 @@ val reduce_to_atomic_ref : env -> evar_map -> global_reference -> types -> types val find_hnf_rectype : - env -> evar_map -> types -> pinductive * constr list + env -> evar_map -> types -> (inductive * EInstance.t) * constr list val contextually : bool -> occurrences * constr_pattern -> (patvar_map -> reduction_function) -> reduction_function diff --git a/pretyping/typeclasses.ml b/pretyping/typeclasses.ml index 478499d91..93c71e6ea 100644 --- a/pretyping/typeclasses.ml +++ b/pretyping/typeclasses.ml @@ -303,6 +303,7 @@ let build_subclasses ~check env sigma glob { hint_priority = pri } = | Some (Forward, info) -> let proj = Option.get proj in let rels = List.map (fun d -> Termops.map_rel_decl EConstr.Unsafe.to_constr d) rels in + let u = EConstr.EInstance.kind sigma u in let body = it_mkLambda_or_LetIn (mkApp (mkConstU (proj,u), projargs)) rels in if check && check_instance env sigma (EConstr.of_constr body) then None else diff --git a/pretyping/typeclasses.mli b/pretyping/typeclasses.mli index 7990b12cd..8d1c0b94c 100644 --- a/pretyping/typeclasses.mli +++ b/pretyping/typeclasses.mli @@ -61,13 +61,13 @@ val class_info : global_reference -> typeclass (** raises a UserError if not a c (** These raise a UserError if not a class. Caution: the typeclass structures is not instantiated w.r.t. the universe instance. This is done separately by typeclass_univ_instance. *) -val dest_class_app : env -> evar_map -> EConstr.constr -> typeclass puniverses * constr list +val dest_class_app : env -> evar_map -> EConstr.constr -> (typeclass * EConstr.EInstance.t) * constr list (** Get the instantiated typeclass structure for a given universe instance. *) val typeclass_univ_instance : typeclass puniverses -> typeclass puniverses (** Just return None if not a class *) -val class_of_constr : evar_map -> EConstr.constr -> (EConstr.rel_context * (typeclass puniverses * constr list)) option +val class_of_constr : evar_map -> EConstr.constr -> (EConstr.rel_context * ((typeclass * EConstr.EInstance.t) * constr list)) option val instance_impl : instance -> global_reference diff --git a/pretyping/typing.ml b/pretyping/typing.ml index d9d64e7eb..c2a030bcd 100644 --- a/pretyping/typing.ml +++ b/pretyping/typing.ml @@ -35,11 +35,13 @@ let meta_type evd mv = let ty = Evd.map_fl EConstr.of_constr ty in meta_instance evd ty -let constant_type_knowing_parameters env sigma cst jl = +let constant_type_knowing_parameters env sigma (cst, u) jl = + let u = Unsafe.to_instance u in let paramstyp = Array.map (fun j -> lazy (EConstr.to_constr sigma j.uj_type)) jl in - EConstr.of_constr (type_of_constant_knowing_parameters_in env cst paramstyp) + EConstr.of_constr (type_of_constant_knowing_parameters_in env (cst, u) paramstyp) let inductive_type_knowing_parameters env sigma (ind,u) jl = + let u = Unsafe.to_instance u in let mspec = lookup_mind_specif env ind in let paramstyp = Array.map (fun j -> lazy (EConstr.to_constr sigma j.uj_type)) jl in EConstr.of_constr (Inductive.type_of_inductive_knowing_parameters env (mspec,u) paramstyp) @@ -140,9 +142,10 @@ let e_type_case_branches env evdref (ind,largs) pj c = (lc, ty) let e_judge_of_case env evdref ci pj cj lfj = - let indspec = + let ((ind, u), spec) = try find_mrectype env !evdref cj.uj_type with Not_found -> error_case_not_inductive env !evdref cj in + let indspec = ((ind, EInstance.kind !evdref u), spec) in let _ = check_case_info env (fst indspec) ci in let (bty,rslty) = e_type_case_branches env evdref indspec pj cj.uj_val in e_check_branch_types env evdref (fst indspec) cj (lfj,bty); @@ -224,6 +227,7 @@ let judge_of_projection env sigma p cj = try find_mrectype env sigma cj.uj_type with Not_found -> error_case_not_inductive env sigma cj in + let u = EInstance.kind sigma u in let ty = EConstr.of_constr (CVars.subst_instance_constr u pb.Declarations.proj_type) in let ty = substl (cj.uj_val :: List.rev args) ty in {uj_val = EConstr.mkProj (p,cj.uj_val); @@ -262,14 +266,17 @@ let rec execute env evdref cstr = | Var id -> judge_of_variable env id - | Const c -> - make_judge cstr (EConstr.of_constr (rename_type_of_constant env c)) + | Const (c, u) -> + let u = EInstance.kind !evdref u in + make_judge cstr (EConstr.of_constr (rename_type_of_constant env (c, u))) - | Ind ind -> - make_judge cstr (EConstr.of_constr (rename_type_of_inductive env ind)) + | Ind (ind, u) -> + let u = EInstance.kind !evdref u in + make_judge cstr (EConstr.of_constr (rename_type_of_inductive env (ind, u))) - | Construct cstruct -> - make_judge cstr (EConstr.of_constr (rename_type_of_constructor env cstruct)) + | Construct (cstruct, u) -> + let u = EInstance.kind !evdref u in + make_judge cstr (EConstr.of_constr (rename_type_of_constructor env (cstruct, u))) | Case (ci,p,c,lf) -> let cj = execute env evdref c in @@ -305,14 +312,14 @@ let rec execute env evdref cstr = let jl = execute_array env evdref args in let j = match EConstr.kind !evdref f with - | Ind ind when Environ.template_polymorphic_pind ind env -> + | Ind (ind, u) when EInstance.is_empty u && Environ.template_polymorphic_ind ind env -> (* Sort-polymorphism of inductive types *) make_judge f - (inductive_type_knowing_parameters env !evdref ind jl) - | Const cst when Environ.template_polymorphic_pconstant cst env -> + (inductive_type_knowing_parameters env !evdref (ind, u) jl) + | Const (cst, u) when EInstance.is_empty u && Environ.template_polymorphic_constant cst env -> (* Sort-polymorphism of inductive types *) make_judge f - (constant_type_knowing_parameters env !evdref cst jl) + (constant_type_knowing_parameters env !evdref (cst, u) jl) | _ -> execute env evdref f in diff --git a/pretyping/unification.ml b/pretyping/unification.ml index 035b0c223..91781a076 100644 --- a/pretyping/unification.ml +++ b/pretyping/unification.ml @@ -535,6 +535,7 @@ let key_of env sigma b flags f = | Const (cst, u) when is_transparent env (ConstKey cst) && (Cpred.mem cst (snd flags.modulo_delta) || Environ.is_projection cst env) -> + let u = EInstance.kind sigma u in Some (IsKey (ConstKey (cst, u))) | Var id when is_transparent env (VarKey id) && Id.Pred.mem id (fst flags.modulo_delta) -> diff --git a/proofs/logic.ml b/proofs/logic.ml index b38a901c2..e6024785d 100644 --- a/proofs/logic.ml +++ b/proofs/logic.ml @@ -505,9 +505,10 @@ and mk_casegoals sigma goal goalacc p c = let (acc',ct,sigma,c') = mk_hdgoals sigma goal goalacc c in let ct = EConstr.of_constr ct in let (acc'',pt,sigma,p') = mk_hdgoals sigma goal acc' p in - let indspec = + let ((ind, u), spec) = try Tacred.find_hnf_rectype env sigma ct with Not_found -> anomaly (Pp.str "mk_casegoals") in + let indspec = ((ind, EConstr.EInstance.kind sigma u), spec) in let (lbrty,conclty) = type_case_branches_with_names env sigma indspec p c in (acc'',lbrty,conclty,sigma,p',c') diff --git a/proofs/tacmach.mli b/proofs/tacmach.mli index 1b05bc7d6..627a8e0e7 100644 --- a/proofs/tacmach.mli +++ b/proofs/tacmach.mli @@ -68,8 +68,8 @@ val pf_whd_all : goal sigma -> constr -> constr val pf_hnf_constr : goal sigma -> constr -> constr val pf_nf : goal sigma -> constr -> constr val pf_nf_betaiota : goal sigma -> constr -> constr -val pf_reduce_to_quantified_ind : goal sigma -> types -> pinductive * types -val pf_reduce_to_atomic_ind : goal sigma -> types -> pinductive * types +val pf_reduce_to_quantified_ind : goal sigma -> types -> (inductive * EInstance.t) * types +val pf_reduce_to_atomic_ind : goal sigma -> types -> (inductive * EInstance.t) * types val pf_compute : goal sigma -> constr -> constr val pf_unfoldn : (occurrences * evaluable_global_reference) list -> goal sigma -> constr -> constr @@ -131,7 +131,7 @@ module New : sig val pf_last_hyp : ('a, 'r) Proofview.Goal.t -> named_declaration val pf_nf_concl : ([ `LZ ], 'r) Proofview.Goal.t -> types - val pf_reduce_to_quantified_ind : ('a, 'r) Proofview.Goal.t -> types -> pinductive * types + val pf_reduce_to_quantified_ind : ('a, 'r) Proofview.Goal.t -> types -> (inductive * EInstance.t) * types val pf_hnf_constr : ('a, 'r) Proofview.Goal.t -> constr -> types val pf_hnf_type_of : ('a, 'r) Proofview.Goal.t -> constr -> types diff --git a/tactics/eauto.ml b/tactics/eauto.ml index e0dff3739..8d1e0e507 100644 --- a/tactics/eauto.ml +++ b/tactics/eauto.ml @@ -472,8 +472,9 @@ let unfold_head env sigma (ids, csts) c = (match Environ.named_body id env with | Some b -> true, EConstr.of_constr b | None -> false, c) - | Const (cst,u as c) when Cset.mem cst csts -> - true, EConstr.of_constr (Environ.constant_value_in env c) + | Const (cst, u) when Cset.mem cst csts -> + let u = EInstance.kind sigma u in + true, EConstr.of_constr (Environ.constant_value_in env (cst, u)) | App (f, args) -> (match aux f with | true, f' -> true, Reductionops.whd_betaiota sigma (mkApp (f', args)) diff --git a/tactics/equality.ml b/tactics/equality.ml index 53b468bff..7ae7446c8 100644 --- a/tactics/equality.ml +++ b/tactics/equality.ml @@ -746,6 +746,7 @@ let find_positions env sigma t1 t2 = let _,rargs2 = List.chop nparams args2 in let (mib,mip) = lookup_mind_specif env ind1 in let params1 = List.map EConstr.Unsafe.to_constr params1 in + let u1 = EInstance.kind sigma u1 in let ctxt = (get_constructor ((ind1,u1),mib,mip,params1) i1).cs_args in let adjust i = CVars.adjust_rel_to_rel_context ctxt (i+1) - 1 in List.flatten @@ -1324,19 +1325,19 @@ let inject_if_homogenous_dependent_pair ty = hd2,ar2 = decompose_app_vect sigma t2 in if not (Termops.is_global sigma (existTconstr()) hd1) then raise Exit; if not (Termops.is_global sigma (existTconstr()) hd2) then raise Exit; - let ind,_ = try pf_apply find_mrectype gl ar1.(0) with Not_found -> raise Exit in + let (ind, _), _ = try pf_apply find_mrectype gl ar1.(0) with Not_found -> raise Exit in (* check if the user has declared the dec principle *) (* and compare the fst arguments of the dep pair *) (* Note: should work even if not an inductive type, but the table only *) (* knows inductive types *) - if not (Ind_tables.check_scheme (!eq_dec_scheme_kind_name()) (fst ind) && + if not (Ind_tables.check_scheme (!eq_dec_scheme_kind_name()) ind && pf_apply is_conv gl ar1.(2) ar2.(2)) then raise Exit; Coqlib.check_required_library ["Coq";"Logic";"Eqdep_dec"]; let new_eq_args = [|pf_unsafe_type_of gl ar1.(3);ar1.(3);ar2.(3)|] in let inj2 = Coqlib.coq_constant "inj_pair2_eq_dec is missing" ["Logic";"Eqdep_dec"] "inj_pair2_eq_dec" in let inj2 = EConstr.of_constr inj2 in - let c, eff = find_scheme (!eq_dec_scheme_kind_name()) (Univ.out_punivs ind) in + let c, eff = find_scheme (!eq_dec_scheme_kind_name()) ind in (* cut with the good equality and prove the requested goal *) tclTHENLIST [Proofview.tclEFFECTS eff; @@ -1783,6 +1784,7 @@ let subst_all ?(flags=default_subst_tactic_flags ()) () = let select_equation_name decl = try let lbeq,u,(_,x,y) = find_eq_data_decompose (NamedDecl.get_type decl) in + let u = EInstance.kind sigma u in let eq = Universes.constr_of_global_univ (lbeq.eq,u) in if flags.only_leibniz then restrict_to_eq_and_identity eq; match EConstr.kind sigma x, EConstr.kind sigma y with @@ -1834,6 +1836,7 @@ let subst_all ?(flags=default_subst_tactic_flags ()) () = let test (_,c) = try let lbeq,u,(_,x,y) = find_eq_data_decompose c in + let u = EInstance.kind sigma u in let eq = Universes.constr_of_global_univ (lbeq.eq,u) in if flags.only_leibniz then restrict_to_eq_and_identity eq; (* J.F.: added to prevent failure on goal containing x=x as an hyp *) diff --git a/tactics/hipattern.ml b/tactics/hipattern.ml index 8e4654c02..851554b83 100644 --- a/tactics/hipattern.ml +++ b/tactics/hipattern.ml @@ -144,6 +144,7 @@ let match_with_tuple sigma t = let t = match_with_one_constructor sigma None false true t in Option.map (fun (hd,l) -> let ind = destInd sigma hd in + let ind = on_snd (fun u -> EInstance.kind sigma u) ind in let (mib,mip) = Global.lookup_pinductive ind in let isrec = mis_is_recursive (fst ind,mib,mip) in (hd,l,isrec)) t @@ -200,8 +201,8 @@ let is_disjunction ?(strict=false) ?(onlybinary=false) sigma t = let match_with_empty_type sigma t = let (hdapp,args) = decompose_app sigma t in match EConstr.kind sigma hdapp with - | Ind ind -> - let (mib,mip) = Global.lookup_pinductive ind in + | Ind (ind, _) -> + let (mib,mip) = Global.lookup_inductive ind in let nconstr = Array.length mip.mind_consnames in if Int.equal nconstr 0 then Some hdapp else None | _ -> None @@ -214,8 +215,8 @@ let is_empty_type sigma t = op2bool (match_with_empty_type sigma t) let match_with_unit_or_eq_type sigma t = let (hdapp,args) = decompose_app sigma t in match EConstr.kind sigma hdapp with - | Ind ind -> - let (mib,mip) = Global.lookup_pinductive ind in + | Ind (ind , _) -> + let (mib,mip) = Global.lookup_inductive ind in let constr_types = mip.mind_nf_lc in let nconstr = Array.length mip.mind_consnames in let zero_args c = Int.equal (nb_prod sigma (EConstr.of_constr c)) mib.mind_nparams in @@ -369,8 +370,8 @@ let is_forall_term sigma c = op2bool (match_with_forall_term sigma c) let match_with_nodep_ind sigma t = let (hdapp,args) = decompose_app sigma t in match EConstr.kind sigma hdapp with - | Ind ind -> - let (mib,mip) = Global.lookup_pinductive ind in + | Ind (ind, _) -> + let (mib,mip) = Global.lookup_inductive ind in if Array.length (mib.mind_packets)>1 then None else let nodep_constr c = has_nodep_prod_after mib.mind_nparams sigma (EConstr.of_constr c) in if Array.for_all nodep_constr mip.mind_nf_lc then @@ -387,8 +388,8 @@ let is_nodep_ind sigma t = op2bool (match_with_nodep_ind sigma t) let match_with_sigma_type sigma t = let (hdapp,args) = decompose_app sigma t in match EConstr.kind sigma hdapp with - | Ind ind -> - let (mib,mip) = Global.lookup_pinductive ind in + | Ind (ind, _) -> + let (mib,mip) = Global.lookup_inductive ind in if Int.equal (Array.length (mib.mind_packets)) 1 && (Int.equal mip.mind_nrealargs 0) && (Int.equal (Array.length mip.mind_consnames)1) && diff --git a/tactics/hipattern.mli b/tactics/hipattern.mli index c46817f50..dd09c3a4d 100644 --- a/tactics/hipattern.mli +++ b/tactics/hipattern.mli @@ -122,19 +122,19 @@ val match_with_equation: (** Match terms [eq A t u], [identity A t u] or [JMeq A t A u] Returns associated lemmas and [A,t,u] or fails PatternMatchingFailure *) val find_eq_data_decompose : ('a, 'r) Proofview.Goal.t -> constr -> - coq_eq_data * Univ.universe_instance * (types * constr * constr) + coq_eq_data * EInstance.t * (types * constr * constr) (** Idem but fails with an error message instead of PatternMatchingFailure *) val find_this_eq_data_decompose : ('a, 'r) Proofview.Goal.t -> constr -> - coq_eq_data * Univ.universe_instance * (types * constr * constr) + coq_eq_data * EInstance.t * (types * constr * constr) (** A variant that returns more informative structure on the equality found *) -val find_eq_data : evar_map -> constr -> coq_eq_data * Univ.universe_instance * equation_kind +val find_eq_data : evar_map -> constr -> coq_eq_data * EInstance.t * equation_kind (** Match a term of the form [(existT A P t p)] Returns associated lemmas and [A,P,t,p] *) val find_sigma_data_decompose : evar_map -> constr -> - coq_sigma_data * (Univ.universe_instance * constr * constr * constr * constr) + coq_sigma_data * (EInstance.t * constr * constr * constr * constr) (** Match a term of the form [{x:A|P}], returns [A] and [P] *) val match_sigma : evar_map -> constr -> constr * constr diff --git a/tactics/tacticals.ml b/tactics/tacticals.ml index a1cd51047..90b7d6581 100644 --- a/tactics/tacticals.ml +++ b/tactics/tacticals.ml @@ -606,6 +606,7 @@ module New = struct isrec allnames tac predicate ind (c, t) = Proofview.Goal.enter { enter = begin fun gl -> let sigma, elim = (mk_elim ind).enter gl in + let ind = on_snd (fun u -> EInstance.kind sigma u) ind in Proofview.tclTHEN (Proofview.Unsafe.tclEVARS sigma) (Proofview.Goal.enter { enter = begin fun gl -> let indclause = mk_clenv_from gl (c, t) in @@ -680,17 +681,19 @@ module New = struct (sigma, EConstr.of_constr c) end } - let gl_make_case_dep ind = { enter = begin fun gl -> + let gl_make_case_dep (ind, u) = { enter = begin fun gl -> let sigma = Sigma.Unsafe.of_evar_map (project gl) in - let Sigma (r, sigma, _) = Indrec.build_case_analysis_scheme (pf_env gl) sigma ind true + let u = EInstance.kind (project gl) u in + let Sigma (r, sigma, _) = Indrec.build_case_analysis_scheme (pf_env gl) sigma (ind, u) true (elimination_sort_of_goal gl) in (Sigma.to_evar_map sigma, EConstr.of_constr r) end } - let gl_make_case_nodep ind = { enter = begin fun gl -> + let gl_make_case_nodep (ind, u) = { enter = begin fun gl -> let sigma = Sigma.Unsafe.of_evar_map (project gl) in - let Sigma (r, sigma, _) = Indrec.build_case_analysis_scheme (pf_env gl) sigma ind false + let u = EInstance.kind (project gl) u in + let Sigma (r, sigma, _) = Indrec.build_case_analysis_scheme (pf_env gl) sigma (ind, u) false (elimination_sort_of_goal gl) in (Sigma.to_evar_map sigma, EConstr.of_constr r) diff --git a/tactics/tacticals.mli b/tactics/tacticals.mli index 5839666a7..3b90ec514 100644 --- a/tactics/tacticals.mli +++ b/tactics/tacticals.mli @@ -124,7 +124,7 @@ val fix_empty_or_and_pattern : int -> delayed_open_constr or_and_intro_pattern_expr -> delayed_open_constr or_and_intro_pattern_expr -val compute_constructor_signatures : rec_flag -> pinductive -> bool list array +val compute_constructor_signatures : rec_flag -> inductive * 'a -> bool list array (** Useful for [as intro_pattern] modifier *) val compute_induction_names : @@ -256,11 +256,11 @@ module New : sig val case_then_using : or_and_intro_pattern option -> (branch_args -> unit Proofview.tactic) -> - constr option -> pinductive -> constr * types -> unit Proofview.tactic + constr option -> inductive * EInstance.t -> constr * types -> unit Proofview.tactic val case_nodep_then_using : or_and_intro_pattern option -> (branch_args -> unit Proofview.tactic) -> - constr option -> pinductive -> constr * types -> unit Proofview.tactic + constr option -> inductive * EInstance.t -> constr * types -> unit Proofview.tactic val elim_on_ba : (branch_assumptions -> unit Proofview.tactic) -> branch_args -> unit Proofview.tactic val case_on_ba : (branch_assumptions -> unit Proofview.tactic) -> branch_args -> unit Proofview.tactic diff --git a/tactics/tactics.ml b/tactics/tactics.ml index 55d6df659..8306ac174 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -1445,6 +1445,7 @@ let general_case_analysis_in_context with_evars clear_flag (c,lbindc) = let t = Retyping.get_type_of env (Sigma.to_evar_map sigma) c in let (mind,_) = reduce_to_quantified_ind env (Sigma.to_evar_map sigma) t in let sort = Tacticals.New.elimination_sort_of_goal gl in + let mind = on_snd (fun u -> EInstance.kind (Sigma.to_evar_map sigma) u) mind in let Sigma (elim, sigma, p) = if occur_term (Sigma.to_evar_map sigma) c concl then build_case_analysis_scheme env sigma mind true sort @@ -1647,6 +1648,7 @@ let descend_in_conjunctions avoid tac (err, info) c = let elim = try DefinedRecord (Recordops.lookup_projections ind) with Not_found -> + let u = EInstance.kind sigma u in let sigma = Sigma.Unsafe.of_evar_map sigma in let Sigma (elim, _, _) = build_case_analysis_scheme env sigma (ind,u) false sort in let elim = EConstr.of_constr elim in @@ -2214,9 +2216,9 @@ let constructor_tac with_evars expctdnumopt i lbind = Array.length (snd (Global.lookup_inductive (fst mind))).mind_consnames in check_number_of_constructors expctdnumopt i nconstr; - let Sigma (cons, sigma, p) = Sigma.fresh_constructor_instance + let Sigma ((cons, u), sigma, p) = Sigma.fresh_constructor_instance (Proofview.Goal.env gl) sigma (fst mind, i) in - let cons = mkConstructU cons in + let cons = mkConstructU (cons, EInstance.make u) in let apply_tac = general_apply true false with_evars None (dloc,(cons,lbind)) in let tac = @@ -4033,24 +4035,25 @@ let compute_elim_signature (evd,(elimc,elimt),ind_type_guess) names_info = let guess_elim isrec dep s hyp0 gl = let tmptyp0 = Tacmach.New.pf_get_hyp_typ hyp0 gl in - let mind,_ = Tacmach.New.pf_reduce_to_quantified_ind gl tmptyp0 in + let (mind, u), _ = Tacmach.New.pf_reduce_to_quantified_ind gl tmptyp0 in let evd, elimc = - if isrec && not (is_nonrec (fst mind)) then find_ind_eliminator (fst mind) s gl + if isrec && not (is_nonrec mind) then find_ind_eliminator mind s gl else let env = Tacmach.New.pf_env gl in let sigma = Sigma.Unsafe.of_evar_map (Tacmach.New.project gl) in + let u = EInstance.kind (Tacmach.New.project gl) u in if use_dependent_propositions_elimination () && dep then - let Sigma (ind, sigma, _) = build_case_analysis_scheme env sigma mind true s in + let Sigma (ind, sigma, _) = build_case_analysis_scheme env sigma (mind, u) true s in let ind = EConstr.of_constr ind in (Sigma.to_evar_map sigma, ind) else - let Sigma (ind, sigma, _) = build_case_analysis_scheme_default env sigma mind s in + let Sigma (ind, sigma, _) = build_case_analysis_scheme_default env sigma (mind, u) s in let ind = EConstr.of_constr ind in (Sigma.to_evar_map sigma, ind) in let elimt = Tacmach.New.pf_unsafe_type_of gl elimc in - evd, ((elimc, NoBindings), elimt), mkIndU mind + evd, ((elimc, NoBindings), elimt), mkIndU (mind, u) let given_elim hyp0 (elimc,lbind as e) gl = let sigma = Tacmach.New.project gl in @@ -4637,9 +4640,10 @@ let case_type t = Proofview.Goal.s_enter { s_enter = begin fun gl -> let sigma = Proofview.Goal.sigma gl in let env = Tacmach.New.pf_env gl in - let (ind,t) = reduce_to_atomic_ind env (Sigma.to_evar_map sigma) t in + let ((ind, u), t) = reduce_to_atomic_ind env (Sigma.to_evar_map sigma) t in + let u = EInstance.kind (Sigma.to_evar_map sigma) u in let s = Tacticals.New.elimination_sort_of_goal gl in - let Sigma (elimc, evd, p) = build_case_analysis_scheme_default env sigma ind s in + let Sigma (elimc, evd, p) = build_case_analysis_scheme_default env sigma (ind, u) s in let elimc = EConstr.of_constr elimc in Sigma (elim_scheme_type elimc t, evd, p) end } diff --git a/vernac/auto_ind_decl.ml b/vernac/auto_ind_decl.ml index 10ac7135b..b9c4c6cc5 100644 --- a/vernac/auto_ind_decl.ml +++ b/vernac/auto_ind_decl.ml @@ -179,9 +179,10 @@ let build_beq_scheme mode kn = *) let compute_A_equality rel_list nlist eqA ndx t = let lifti = ndx in + let sigma = Evd.empty (** FIXME *) in let rec aux c = let (c,a) = Reductionops.whd_betaiota_stack Evd.empty c in - match EConstr.kind Evd.empty (** FIXME *) c with + match EConstr.kind sigma c with | Rel x -> mkRel (x-nlist+ndx), Safe_typing.empty_private_constants | Var x -> let eid = id_of_string ("eq_"^(string_of_id x)) in @@ -215,9 +216,10 @@ let build_beq_scheme mode kn = | Prod _ -> raise InductiveWithProduct | Lambda _-> raise (EqUnknown "abstraction") | LetIn _ -> raise (EqUnknown "let-in") - | Const kn -> - (match Environ.constant_opt_value_in env kn with - | None -> raise (ParameterWithoutEquality (ConstRef (fst kn))) + | Const (kn, u) -> + let u = EConstr.EInstance.kind sigma u in + (match Environ.constant_opt_value_in env (kn, u) with + | None -> raise (ParameterWithoutEquality (ConstRef kn)) | Some c -> aux (EConstr.applist (EConstr.of_constr c,a))) | Proj _ -> raise (EqUnknown "projection") | Construct _ -> raise (EqUnknown "constructor") @@ -373,7 +375,7 @@ let do_replace_lb mode lb_scheme_key aavoid narg p q = let u,v = destruct_ind sigma type_of_pq in let lb_type_of_p = try - let c, eff = find_scheme ~mode lb_scheme_key (out_punivs u) (*FIXME*) in + let c, eff = find_scheme ~mode lb_scheme_key (fst u) (*FIXME*) in Proofview.tclUNIT (mkConst c, eff) with Not_found -> (* spiwack: the format of this error message should probably @@ -445,7 +447,7 @@ let do_replace_bl mode bl_scheme_key (ind,u as indu) aavoid narg lft rgt = else ( let bl_t1, eff = try - let c, eff = find_scheme bl_scheme_key (out_punivs u) (*FIXME*) in + let c, eff = find_scheme bl_scheme_key (fst u) (*FIXME*) in mkConst c, eff with Not_found -> (* spiwack: the format of this error message should probably @@ -485,13 +487,13 @@ let do_replace_bl mode bl_scheme_key (ind,u as indu) aavoid narg lft rgt = begin try Proofview.tclUNIT (destApp sigma rgt) with DestKO -> Tacticals.New.tclZEROMSG (str "replace failed.") end >>= fun (ind2,ca2) -> - begin try Proofview.tclUNIT (out_punivs (destInd sigma ind1)) + begin try Proofview.tclUNIT (fst (destInd sigma ind1)) with DestKO -> begin try Proofview.tclUNIT (fst (fst (destConstruct sigma ind1))) with DestKO -> Tacticals.New.tclZEROMSG (str "The expected type is an inductive one.") end end >>= fun (sp1,i1) -> - begin try Proofview.tclUNIT (out_punivs (destInd sigma ind2)) + begin try Proofview.tclUNIT (fst (destInd sigma ind2)) with DestKO -> begin try Proofview.tclUNIT (fst (fst (destConstruct sigma ind2))) with DestKO -> Tacticals.New.tclZEROMSG (str "The expected type is an inductive one.") @@ -664,7 +666,7 @@ let make_bl_scheme mode mind = let side_eff = side_effect_of_mode mode in let bl_goal = EConstr.of_constr bl_goal in let (ans, _, ctx) = Pfedit.build_by_tactic ~side_eff (Global.env()) ctx bl_goal - (compute_bl_tact mode (!bl_scheme_kind_aux()) (ind, Univ.Instance.empty) lnamesparrec nparrec) + (compute_bl_tact mode (!bl_scheme_kind_aux()) (ind, EConstr.EInstance.empty) lnamesparrec nparrec) in ([|ans|], ctx), eff diff --git a/vernac/classes.ml b/vernac/classes.ml index 415b98f2d..c3300a361 100644 --- a/vernac/classes.ml +++ b/vernac/classes.ml @@ -166,8 +166,9 @@ let new_instance ?(abstract=false) ?(global=false) ?(refine= !refine_instance) p let imps = imps @ Impargs.lift_implicits len imps' in let ctx', c = decompose_prod_assum c' in let ctx'' = ctx' @ ctx in - let k, args = Typeclasses.dest_class_app (push_rel_context ctx'' env) !evars (EConstr.of_constr c) in - let cl, u = Typeclasses.typeclass_univ_instance k in + let (k, u), args = Typeclasses.dest_class_app (push_rel_context ctx'' env) !evars (EConstr.of_constr c) in + let u = EConstr.EInstance.kind !evars u in + let cl, u = Typeclasses.typeclass_univ_instance (k, u) in let _, args = List.fold_right (fun decl (args, args') -> match decl with -- cgit v1.2.3 From b824d8ad00001f6c41d0fc8bbf528dccb937c887 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sat, 1 Apr 2017 20:25:05 +0200 Subject: Restore a fast path in EConstr instance normalization. --- engine/eConstr.ml | 4 +++- 1 file changed, 3 insertions(+), 1 deletion(-) diff --git a/engine/eConstr.ml b/engine/eConstr.ml index 28e9ffdb2..bb9075e74 100644 --- a/engine/eConstr.ml +++ b/engine/eConstr.ml @@ -63,7 +63,9 @@ module EInstance = struct type t = Univ.Instance.t let make i = i - let kind sigma i = Evd.normalize_universe_instance sigma 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 -- cgit v1.2.3 From 2794b3c91bbbef115303b40f2e494ad97467dc9e Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Wed, 5 Apr 2017 14:05:42 +0200 Subject: Removing a normalization hotspot from EConstr. It was not necessary to normalize a term just to check whether it was a global reference. The hotspot appeared in mathcomp. --- pretyping/evarconv.ml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/pretyping/evarconv.ml b/pretyping/evarconv.ml index 9c9350ab1..44b771283 100644 --- a/pretyping/evarconv.ml +++ b/pretyping/evarconv.ml @@ -154,7 +154,7 @@ let check_conv_record env sigma (t1,sk1) (t2,sk2) = lookup_canonical_conversion (proji, Sort_cs (family_of_sort s)),[] | _ -> - let c2 = global_of_constr (EConstr.to_constr sigma t2) in + let (c2, _) = Termops.global_of_constr sigma t2 in lookup_canonical_conversion (proji, Const_cs c2),sk2 with Not_found -> let (c, cs) = lookup_canonical_conversion (proji,Default_cs) in -- cgit v1.2.3 From d6175b9980808ff91f1299ca26a9a49a117169ca Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Thu, 6 Apr 2017 17:34:23 +0200 Subject: Fix a normalization hotspot in computation of constr keys. Getting a key only needs to observe the root of a term. This hotspot was observed in HoTT. --- library/keys.ml | 4 ++-- library/keys.mli | 2 +- plugins/ltac/extratactics.ml4 | 12 ++++++++---- pretyping/unification.ml | 4 ++-- 4 files changed, 13 insertions(+), 9 deletions(-) diff --git a/library/keys.ml b/library/keys.ml index 057dc3b65..c9e325ee5 100644 --- a/library/keys.ml +++ b/library/keys.ml @@ -114,11 +114,11 @@ let inKeys : key_obj -> obj = let declare_equiv_keys ref ref' = Lib.add_anonymous_leaf (inKeys (ref,ref')) -let constr_key c = +let constr_key kind c = let open Globnames in try let rec aux k = - match kind_of_term k with + match kind k with | Const (c, _) -> KGlob (ConstRef c) | Ind (i, u) -> KGlob (IndRef i) | Construct (c,u) -> KGlob (ConstructRef c) diff --git a/library/keys.mli b/library/keys.mli index 69668590d..6abac4de4 100644 --- a/library/keys.mli +++ b/library/keys.mli @@ -16,7 +16,7 @@ val declare_equiv_keys : key -> key -> unit val equiv_keys : key -> key -> bool (** Check equivalence of keys. *) -val constr_key : Term.constr -> key option +val constr_key : ('a -> ('a, 't, 'u, 'i) Constr.kind_of_term) -> 'a -> key option (** Compute the head key of a term. *) val pr_keys : (global_reference -> Pp.std_ppcmds) -> Pp.std_ppcmds diff --git a/plugins/ltac/extratactics.ml4 b/plugins/ltac/extratactics.ml4 index b34afd51b..1ec52718a 100644 --- a/plugins/ltac/extratactics.ml4 +++ b/plugins/ltac/extratactics.ml4 @@ -1087,10 +1087,14 @@ END (** library/keys *) VERNAC COMMAND EXTEND Declare_keys CLASSIFIED AS SIDEFF -| [ "Declare" "Equivalent" "Keys" constr(c) constr(c') ] -> [ - let it c = EConstr.Unsafe.to_constr (snd (Constrintern.interp_open_constr (Global.env ()) Evd.empty c)) in - let k1 = Keys.constr_key (it c) in - let k2 = Keys.constr_key (it c') in +| [ "Declare" "Equivalent" "Keys" constr(c) constr(c') ] -> [ + let get_key c = + let (evd, c) = Constrintern.interp_open_constr (Global.env ()) Evd.empty c in + let kind c = EConstr.kind evd c in + Keys.constr_key kind c + in + let k1 = get_key c in + let k2 = get_key c' in match k1, k2 with | Some k1, Some k2 -> Keys.declare_equiv_keys k1 k2 | _ -> () ] diff --git a/pretyping/unification.ml b/pretyping/unification.ml index 91781a076..eb90dfbdb 100644 --- a/pretyping/unification.ml +++ b/pretyping/unification.ml @@ -1739,7 +1739,7 @@ let keyed_unify env evd kop = | None -> fun _ -> true | Some kop -> fun cl -> - let kc = Keys.constr_key (EConstr.to_constr evd cl) in + let kc = Keys.constr_key (fun c -> EConstr.kind evd c) cl in match kc with | None -> false | Some kc -> Keys.equiv_keys kop kc @@ -1749,7 +1749,7 @@ let keyed_unify env evd kop = Fails if no match is found *) let w_unify_to_subterm env evd ?(flags=default_unify_flags ()) (op,cl) = let bestexn = ref None in - let kop = Keys.constr_key (EConstr.to_constr evd op) in + let kop = Keys.constr_key (fun c -> EConstr.kind evd c) op in let rec matchrec cl = let cl = strip_outer_cast evd cl in (try -- cgit v1.2.3 From 2ba1351849550c50e0fe2c7fc7f63758c10fb14a Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 7 Apr 2017 16:31:25 +0200 Subject: Fix an unhandled exception in Omega. --- plugins/omega/coq_omega.ml | 1 + 1 file changed, 1 insertion(+) diff --git a/plugins/omega/coq_omega.ml b/plugins/omega/coq_omega.ml index a2016cb03..7780de712 100644 --- a/plugins/omega/coq_omega.ml +++ b/plugins/omega/coq_omega.ml @@ -1884,6 +1884,7 @@ let destructure_goal = (mkApp (Lazy.force coq_dec_not_not, [| t; dec; mkNewMeta () |])))) intro with Undecidable -> Tactics.elim_type (Lazy.force coq_False) + | e when Proofview.V82.catchable_exception e -> Proofview.tclZERO e in Tacticals.New.tclTHEN goal_tac destructure_hyps in -- cgit v1.2.3 From 9d1230d484a2cf519f9cd76dc0f37815f3c6339b Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sat, 8 Apr 2017 01:42:19 +0200 Subject: Fix a heuristic used by legacy typeclass resolution. The evarmap used by the heuristic could contain resolved evars, which could lead to a failure of backtracking in the EConstr branch. This is experimental and may be to costly. --- tactics/class_tactics.ml | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/tactics/class_tactics.ml b/tactics/class_tactics.ml index 7eadde78d..ea1966093 100644 --- a/tactics/class_tactics.ml +++ b/tactics/class_tactics.ml @@ -787,10 +787,10 @@ module V85 = struct let fk' = (fun e -> let do_backtrack = - if unique then occur_existential s' concl + if unique then occur_existential tacgl.sigma concl else if info.unique then true else if List.is_empty gls' then - needs_backtrack env s' info.is_evar concl + needs_backtrack env tacgl.sigma info.is_evar concl else true in let e' = match foundone with None -> e -- cgit v1.2.3 From 1d3500af261d01ec552424c46311d49fe61c1300 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Mon, 10 Apr 2017 16:35:19 +0200 Subject: Documenting the changes introduced by the EConstr branch. --- CHANGES | 11 +++++++++++ 1 file changed, 11 insertions(+) diff --git a/CHANGES b/CHANGES index 240d71ec0..f42239c79 100644 --- a/CHANGES +++ b/CHANGES @@ -6,6 +6,17 @@ Tactics - New tactic "extensionality in H" which applies (possibly dependent) functional extensionality in H supposed to be a quantified equality until giving a bare equality. +- New representation of terms that statically ensure stability by + evar-expansion. This has several consequences. + * In terms of performance, this adds a cost to every term destructuration, + but at the same time most eager evar normalizations were removed, which + couterbalances this drawback and even sometimes outperforms the old + implementation. For instance, many operations that would require O(n) + normalization of the term are now O(1) in tactics. YMMV. + * This triggers small changes in unification, which was not evar-insensitive. + Most notably, the new implementation recognizes Miller patterns that were + missed before because of a missing normalization step. Hopefully this should + be fairly uncommon. Standard Library -- cgit v1.2.3 From d331543b7b759bb97ea1ece32501cd0149627e9f Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Mon, 10 Apr 2017 16:37:11 +0200 Subject: Adding a test for the correctness of normalization in legacy typeclasses. This is a test for commit 9d1230d484a2cf519f9cd76dc0f37815f3c6339b. --- test-suite/success/old_typeclass.v | 13 +++++++++++++ 1 file changed, 13 insertions(+) create mode 100644 test-suite/success/old_typeclass.v diff --git a/test-suite/success/old_typeclass.v b/test-suite/success/old_typeclass.v new file mode 100644 index 000000000..01e35810b --- /dev/null +++ b/test-suite/success/old_typeclass.v @@ -0,0 +1,13 @@ +Require Import Setoid Coq.Classes.Morphisms. +Set Typeclasses Legacy Resolution. + +Declare Instance and_Proper_eq: Proper (Logic.eq ==> Logic.eq ==> Logic.eq) and. + +Axiom In : Prop. +Axiom union_spec : In <-> True. + +Lemma foo : In /\ True. +Proof. +progress rewrite union_spec. +repeat constructor. +Qed. -- cgit v1.2.3 From 2e6a89238dc7197057d0da80a16f4b4b1e41bfd8 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Mon, 10 Apr 2017 18:35:18 +0200 Subject: Adding a test for 'rewrite in *' when an evar is solved by side-effect. --- test-suite/success/rewrite_evar.v | 8 ++++++++ 1 file changed, 8 insertions(+) create mode 100644 test-suite/success/rewrite_evar.v diff --git a/test-suite/success/rewrite_evar.v b/test-suite/success/rewrite_evar.v new file mode 100644 index 000000000..f7ad261cb --- /dev/null +++ b/test-suite/success/rewrite_evar.v @@ -0,0 +1,8 @@ +Require Import Coq.Setoids.Setoid. + +Goal forall (T2 MT1 MT2 : Type) (x : T2) (M2 m2 : MT2) (M1 m1 : MT1) (F : T2 -> MT1 -> MT2 -> Prop), + (forall (defaultB : T2) (m3 : MT1) (m4 : MT2), F defaultB m3 m4 <-> True) -> F x M1 M2 -> F x m1 m2. + intros ????????? H' H. + rewrite (H' _) in *. + (** The above rewrite should also rewrite in H. *) + Fail progress rewrite H' in H. -- cgit v1.2.3