From 27c93465dcf0d5233c1195d1649f5e699ed54a90 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sun, 21 Aug 2016 20:55:23 +0200 Subject: Packing the named_context_val in a proper record and marking it private. --- kernel/environ.ml | 108 +++++++++++++++++++++++------------------------------- 1 file changed, 46 insertions(+), 62 deletions(-) (limited to 'kernel/environ.ml') diff --git a/kernel/environ.ml b/kernel/environ.ml index 7351a87d4..cfdd93284 100644 --- a/kernel/environ.ml +++ b/kernel/environ.ml @@ -56,14 +56,14 @@ let type_in_type env = not (typing_flags env).check_universes let deactivated_guard env = not (typing_flags env).check_guarded let universes env = env.env_stratification.env_universes -let named_context env = env.env_named_context -let named_context_val env = env.env_named_context,env.env_named_vals +let named_context env = env.env_named_context.env_named_ctx +let named_context_val env = env.env_named_context let rel_context env = env.env_rel_context let opaque_tables env = env.indirect_pterms let set_opaque_tables env indirect_pterms = { env with indirect_pterms } let empty_context env = - match env.env_rel_context, env.env_named_context with + match env.env_rel_context, env.env_named_context.env_named_ctx with | [], [] -> true | _ -> false @@ -99,14 +99,18 @@ let fold_rel_context f env ~init = (* Named context *) -let named_context_of_val = fst -let named_vals_of_val = snd +let named_context_of_val c = c.env_named_ctx +let named_vals_of_val c = c.env_named_val (* [map_named_val f ctxt] apply [f] to the body and the type of each declarations. *** /!\ *** [f t] should be convertible with t *) -let map_named_val f = - on_fst (Context.Named.map f) +let rec map_named_val f ctx = match match_named_context_val ctx with +| None -> empty_named_context_val +| Some (d, v, ctx) -> + let d = Context.Named.Declaration.map_constr f d in + let ctx = map_named_val f ctx in + push_named_context_val_val d v ctx let empty_named_context = Context.Named.empty @@ -118,8 +122,8 @@ let val_of_named_context ctxt = List.fold_right push_named_context_val ctxt empty_named_context_val -let lookup_named id env = Context.Named.lookup id env.env_named_context -let lookup_named_val id (ctxt,_) = Context.Named.lookup id ctxt +let lookup_named id env = Context.Named.lookup id env.env_named_context.env_named_ctx +let lookup_named_val id ctxt = Context.Named.lookup id ctxt.env_named_ctx let eq_named_context_val c1 c2 = c1 == c2 || Context.Named.equal (named_context_of_val c1) (named_context_of_val c2) @@ -139,10 +143,9 @@ let evaluable_named id env = | Some _ -> true | _ -> false -let reset_with_named_context (ctxt,ctxtv) env = +let reset_with_named_context ctxt env = { env with env_named_context = ctxt; - env_named_vals = ctxtv; env_rel_context = Context.Rel.empty; env_rel_val = []; env_nb_rel = 0 } @@ -157,11 +160,11 @@ let pop_rel_context n env = let fold_named_context f env ~init = let rec fold_right env = - match env.env_named_context with - | [] -> init - | d::ctxt -> + match match_named_context_val env.env_named_context with + | None -> init + | Some (d, v, rem) -> let env = - reset_with_named_context (ctxt,List.tl env.env_named_vals) env in + reset_with_named_context rem env in f env d (fold_right env) in fold_right env @@ -493,66 +496,47 @@ let compile_constant_body = Cbytegen.compile_constant_body false exception Hyp_not_found -let apply_to_hyp (ctxt,vals) id f = - let rec aux rtail ctxt vals = - match ctxt, vals with - | d::ctxt, v::vals -> +let apply_to_hyp ctxt id f = + let rec aux rtail ctxt = + match match_named_context_val ctxt with + | Some (d, v, ctxt) -> if Id.equal (get_id d) id then - (f ctxt d rtail)::ctxt, v::vals + push_named_context_val_val (f ctxt.env_named_ctx d rtail) v ctxt else - let ctxt',vals' = aux (d::rtail) ctxt vals in - d::ctxt', v::vals' - | [],[] -> raise Hyp_not_found - | _, _ -> assert false - in aux [] ctxt vals - -let apply_to_hyp_and_dependent_on (ctxt,vals) id f g = - let rec aux ctxt vals = - match ctxt,vals with - | d::ctxt, v::vals -> + let ctxt' = aux (d::rtail) ctxt in + push_named_context_val_val d v ctxt' + | None -> raise Hyp_not_found + in aux [] ctxt + +let apply_to_hyp_and_dependent_on ctxt id f g = + let rec aux sign = + match match_named_context_val sign with + | Some (d, v, sign) -> if Id.equal (get_id d) id then - let sign = ctxt,vals in push_named_context_val (f d sign) sign else - let (ctxt,vals as sign) = aux ctxt vals in + let sign = aux sign in push_named_context_val (g d sign) sign - | [],[] -> raise Hyp_not_found - | _,_ -> assert false - in aux ctxt vals - -let insert_after_hyp (ctxt,vals) id d check = - let rec aux ctxt vals = - match ctxt, vals with - | decl::ctxt', v::vals' -> - if Id.equal (get_id decl) id then begin - check ctxt; - push_named_context_val d (ctxt,vals) - end else - let ctxt,vals = aux ctxt vals in - d::ctxt, v::vals - | [],[] -> raise Hyp_not_found - | _, _ -> assert false - in aux ctxt vals - + | None -> raise Hyp_not_found + in aux ctxt (* To be used in Logic.clear_hyps *) -let remove_hyps ids check_context check_value (ctxt, vals) = - let rec remove_hyps ctxt vals = match ctxt, vals with - | [], [] -> ([], []), false - | d :: rctxt, (nid, v) :: rvals -> - let (ans, seen) = remove_hyps rctxt rvals in +let remove_hyps ids check_context check_value ctxt = + let rec remove_hyps ctxt = match match_named_context_val ctxt with + | None -> empty_named_context_val, false + | Some (d, v, rctxt) -> + let (ans, seen) = remove_hyps rctxt in if Id.Set.mem (get_id d) ids then (ans, true) - else if not seen then (ctxt, vals), false + else if not seen then ctxt, false else - let (rctxt', rvals') = ans in + let rctxt' = ans in let d' = check_context d in let v' = check_value v in - if d == d' && v == v' && rctxt == rctxt' && rvals == rvals' then - (ctxt, vals), true - else (d' :: rctxt', (nid, v') :: rvals'), true - | _ -> assert false + if d == d' && v == v' && rctxt == rctxt' then + ctxt, true + else push_named_context_val_val d' v' rctxt', true in - fst (remove_hyps ctxt vals) + fst (remove_hyps ctxt) (*spiwack: the following functions assemble the pieces of the retroknowledge note that the "consistent" register function is available in the module -- cgit v1.2.3 From e046e9b7e35cbe2d419099907cdcc0909b59d52c Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sun, 21 Aug 2016 21:43:16 +0200 Subject: Tentative fast-access named env --- kernel/environ.ml | 4 ++-- kernel/pre_env.ml | 9 +++++++-- kernel/pre_env.mli | 1 + 3 files changed, 10 insertions(+), 4 deletions(-) (limited to 'kernel/environ.ml') diff --git a/kernel/environ.ml b/kernel/environ.ml index cfdd93284..be3b4977e 100644 --- a/kernel/environ.ml +++ b/kernel/environ.ml @@ -122,8 +122,8 @@ let val_of_named_context ctxt = List.fold_right push_named_context_val ctxt empty_named_context_val -let lookup_named id env = Context.Named.lookup id env.env_named_context.env_named_ctx -let lookup_named_val id ctxt = Context.Named.lookup id ctxt.env_named_ctx +let lookup_named id env = fst (Id.Map.find id env.env_named_context.env_named_map) +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) diff --git a/kernel/pre_env.ml b/kernel/pre_env.ml index 47d003fe9..62e5cd366 100644 --- a/kernel/pre_env.ml +++ b/kernel/pre_env.ml @@ -66,6 +66,7 @@ type named_vals = (Id.t * lazy_val) list type named_context_val = { env_named_ctx : Context.Named.t; env_named_val : named_vals; + env_named_map : (Context.Named.Declaration.t * lazy_val) Id.Map.t; } type env = { @@ -84,6 +85,7 @@ type env = { let empty_named_context_val = { env_named_ctx = []; env_named_val = []; + env_named_map = Id.Map.empty; } let empty_env = { @@ -130,9 +132,11 @@ let env_of_rel n env = (* Named context *) let push_named_context_val_val d rval ctxt = + assert (not (Id.Map.mem (get_id d) ctxt.env_named_map)); { env_named_ctx = Context.Named.add d ctxt.env_named_ctx; env_named_val = (get_id d, rval) :: ctxt.env_named_val; + env_named_map = Id.Map.add (get_id d) (d, rval) ctxt.env_named_map; } let push_named_context_val d ctxt = @@ -141,7 +145,8 @@ let push_named_context_val d ctxt = let match_named_context_val c = match c.env_named_ctx, c.env_named_val with | [], [] -> None | decl :: ctx, (_, v) :: vls -> - let cval = { env_named_ctx = ctx; env_named_val = vls } in + let map = Id.Map.remove (get_id decl) c.env_named_map in + let cval = { env_named_ctx = ctx; env_named_val = vls; env_named_map = map } in Some (decl, v, cval) | _ -> assert false @@ -161,7 +166,7 @@ let push_named d env = } let lookup_named_val id env = - snd(List.find (fun (id',_) -> Id.equal id id') env.env_named_context.env_named_val) + snd(Id.Map.find id env.env_named_context.env_named_map) (* Warning all the names should be different *) let env_of_named id env = env diff --git a/kernel/pre_env.mli b/kernel/pre_env.mli index 928a8009e..f916c399f 100644 --- a/kernel/pre_env.mli +++ b/kernel/pre_env.mli @@ -45,6 +45,7 @@ type named_vals = (Id.t * lazy_val) list type named_context_val = private { env_named_ctx : Context.Named.t; env_named_val : named_vals; + env_named_map : (Context.Named.Declaration.t * lazy_val) Id.Map.t; } type env = { -- cgit v1.2.3 From 37195c027472b7f0110249b28356271e713fee6f Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Mon, 22 Aug 2016 14:39:10 +0200 Subject: More efficient implementation of map_named_val. --- kernel/environ.ml | 7 +------ kernel/pre_env.ml | 15 ++++++++++++++- kernel/pre_env.mli | 3 +++ 3 files changed, 18 insertions(+), 7 deletions(-) (limited to 'kernel/environ.ml') diff --git a/kernel/environ.ml b/kernel/environ.ml index be3b4977e..f4a3312ef 100644 --- a/kernel/environ.ml +++ b/kernel/environ.ml @@ -105,12 +105,7 @@ let named_vals_of_val c = c.env_named_val (* [map_named_val f ctxt] apply [f] to the body and the type of each declarations. *** /!\ *** [f t] should be convertible with t *) -let rec map_named_val f ctx = match match_named_context_val ctx with -| None -> empty_named_context_val -| Some (d, v, ctx) -> - let d = Context.Named.Declaration.map_constr f d in - let ctx = map_named_val f ctx in - push_named_context_val_val d v ctx +let map_named_val = map_named_val let empty_named_context = Context.Named.empty diff --git a/kernel/pre_env.ml b/kernel/pre_env.ml index 62e5cd366..defc0afc8 100644 --- a/kernel/pre_env.ml +++ b/kernel/pre_env.ml @@ -132,7 +132,7 @@ let env_of_rel n env = (* Named context *) let push_named_context_val_val d rval ctxt = - assert (not (Id.Map.mem (get_id d) ctxt.env_named_map)); +(* assert (not (Id.Map.mem (get_id d) ctxt.env_named_map)); *) { env_named_ctx = Context.Named.add d ctxt.env_named_ctx; env_named_val = (get_id d, rval) :: ctxt.env_named_val; @@ -150,6 +150,19 @@ let match_named_context_val c = match c.env_named_ctx, c.env_named_val with Some (decl, v, cval) | _ -> assert false +let map_named_val f ctxt = + let open Context.Named.Declaration in + let fold accu d = + let d' = map_constr f d in + let accu = + if d == d' then accu + else Id.Map.modify (get_id d) (fun _ (_, v) -> (d', v)) accu + in + (accu, d') + in + let map, ctx = List.fold_map fold ctxt.env_named_map ctxt.env_named_ctx in + { ctxt with env_named_ctx = ctx; env_named_map = map } + let push_named d env = (* if not (env.env_rel_context = []) then raise (ASSERT env.env_rel_context); assert (env.env_rel_context = []); *) diff --git a/kernel/pre_env.mli b/kernel/pre_env.mli index f916c399f..4390177da 100644 --- a/kernel/pre_env.mli +++ b/kernel/pre_env.mli @@ -80,6 +80,9 @@ val push_named_context_val_val : Context.Named.Declaration.t -> lazy_val -> named_context_val -> named_context_val val match_named_context_val : named_context_val -> (Context.Named.Declaration.t * lazy_val * named_context_val) option +val map_named_val : + (constr -> constr) -> named_context_val -> named_context_val + val push_named : Context.Named.Declaration.t -> env -> env val lookup_named_val : Id.t -> env -> lazy_val val env_of_named : Id.t -> env -> env -- cgit v1.2.3 From 1888527bb43d6a8c801565af3e6376c91769fbc1 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 9 Sep 2016 11:51:51 +0200 Subject: Removing the now useless field env_named_val from named_context_val. This field was only used by the VM before, but since the previous patches, this part of the code relies on the map instead. --- kernel/environ.ml | 1 - kernel/environ.mli | 1 - kernel/pre_env.ml | 17 ++++++----------- kernel/pre_env.mli | 3 --- 4 files changed, 6 insertions(+), 16 deletions(-) (limited to 'kernel/environ.ml') diff --git a/kernel/environ.ml b/kernel/environ.ml index f4a3312ef..96e35610c 100644 --- a/kernel/environ.ml +++ b/kernel/environ.ml @@ -100,7 +100,6 @@ let fold_rel_context f env ~init = (* Named context *) let named_context_of_val c = c.env_named_ctx -let named_vals_of_val c = c.env_named_val (* [map_named_val f ctxt] apply [f] to the body and the type of each declarations. diff --git a/kernel/environ.mli b/kernel/environ.mli index cb4a94d45..235ba2fd1 100644 --- a/kernel/environ.mli +++ b/kernel/environ.mli @@ -78,7 +78,6 @@ val fold_rel_context : (** {5 Context of variables (section variables and goal assumptions) } *) val named_context_of_val : named_context_val -> Context.Named.t -val named_vals_of_val : named_context_val -> Pre_env.named_vals val val_of_named_context : Context.Named.t -> named_context_val val empty_named_context_val : named_context_val diff --git a/kernel/pre_env.ml b/kernel/pre_env.ml index defc0afc8..104fe59d1 100644 --- a/kernel/pre_env.ml +++ b/kernel/pre_env.ml @@ -61,11 +61,8 @@ let force_lazy_val vk = match !vk with let dummy_lazy_val () = ref VKnone let build_lazy_val vk key = vk := VKvalue (CEphemeron.create key) -type named_vals = (Id.t * lazy_val) list - type named_context_val = { env_named_ctx : Context.Named.t; - env_named_val : named_vals; env_named_map : (Context.Named.Declaration.t * lazy_val) Id.Map.t; } @@ -84,7 +81,6 @@ type env = { let empty_named_context_val = { env_named_ctx = []; - env_named_val = []; env_named_map = Id.Map.empty; } @@ -135,20 +131,19 @@ let push_named_context_val_val d rval ctxt = (* assert (not (Id.Map.mem (get_id d) ctxt.env_named_map)); *) { env_named_ctx = Context.Named.add d ctxt.env_named_ctx; - env_named_val = (get_id d, rval) :: ctxt.env_named_val; env_named_map = Id.Map.add (get_id d) (d, rval) ctxt.env_named_map; } let push_named_context_val d ctxt = push_named_context_val_val d (ref VKnone) ctxt -let match_named_context_val c = match c.env_named_ctx, c.env_named_val with -| [], [] -> None -| decl :: ctx, (_, v) :: vls -> +let match_named_context_val c = match c.env_named_ctx with +| [] -> None +| decl :: ctx -> + let (_, v) = Id.Map.find (get_id decl) c.env_named_map in let map = Id.Map.remove (get_id decl) c.env_named_map in - let cval = { env_named_ctx = ctx; env_named_val = vls; env_named_map = map } in + let cval = { env_named_ctx = ctx; env_named_map = map } in Some (decl, v, cval) -| _ -> assert false let map_named_val f ctxt = let open Context.Named.Declaration in @@ -161,7 +156,7 @@ let map_named_val f ctxt = (accu, d') in let map, ctx = List.fold_map fold ctxt.env_named_map ctxt.env_named_ctx in - { ctxt with env_named_ctx = ctx; env_named_map = map } + { env_named_ctx = ctx; env_named_map = map } let push_named d env = (* if not (env.env_rel_context = []) then raise (ASSERT env.env_rel_context); diff --git a/kernel/pre_env.mli b/kernel/pre_env.mli index 4390177da..1e4eff3d5 100644 --- a/kernel/pre_env.mli +++ b/kernel/pre_env.mli @@ -40,11 +40,8 @@ val force_lazy_val : lazy_val -> (values * Id.Set.t) option val dummy_lazy_val : unit -> lazy_val val build_lazy_val : lazy_val -> (values * Id.Set.t) -> unit -type named_vals = (Id.t * lazy_val) list - type named_context_val = private { env_named_ctx : Context.Named.t; - env_named_val : named_vals; env_named_map : (Context.Named.Declaration.t * lazy_val) Id.Map.t; } -- cgit v1.2.3 From d55818c7da468ce1c7c9644cb63f68f7561a17bc Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Fri, 9 Sep 2016 13:29:28 +0200 Subject: Tracking careless uses of slow name lookup. --- engine/termops.ml | 7 ++----- engine/termops.mli | 2 +- kernel/environ.ml | 2 +- kernel/nativecode.ml | 2 +- kernel/pre_env.ml | 3 +++ kernel/pre_env.mli | 1 + pretyping/unification.ml | 2 +- proofs/logic.ml | 2 +- proofs/tacmach.ml | 4 ++-- tactics/leminv.ml | 4 ++-- tactics/tactics.ml | 24 ++++++++++++------------ 11 files changed, 27 insertions(+), 26 deletions(-) (limited to 'kernel/environ.ml') diff --git a/engine/termops.ml b/engine/termops.ml index a047bf53c..17e56ec31 100644 --- a/engine/termops.ml +++ b/engine/termops.ml @@ -975,11 +975,8 @@ let smash_rel_context sign = let fold_named_context_both_sides f l ~init = List.fold_right_and_left f l init -let rec mem_named_context id ctxt = - match ctxt with - | decl :: _ when Id.equal id (NamedDecl.get_id decl) -> true - | _ :: sign -> mem_named_context id sign - | [] -> false +let mem_named_context_val id ctxt = + try Environ.lookup_named_val id ctxt; true with Not_found -> false let compact_named_context_reverse sign = let compact l decl = diff --git a/engine/termops.mli b/engine/termops.mli index 5d85088f8..0a7ac1f26 100644 --- a/engine/termops.mli +++ b/engine/termops.mli @@ -237,7 +237,7 @@ val map_rel_context_with_binders : 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 : Id.t -> Context.Named.t -> bool +val mem_named_context_val : Id.t -> named_context_val -> bool val compact_named_context : Context.Named.t -> Context.NamedList.t val compact_named_context_reverse : Context.Named.t -> Context.NamedList.t diff --git a/kernel/environ.ml b/kernel/environ.ml index 96e35610c..54898320d 100644 --- a/kernel/environ.ml +++ b/kernel/environ.ml @@ -116,7 +116,7 @@ let val_of_named_context ctxt = List.fold_right push_named_context_val ctxt empty_named_context_val -let lookup_named id env = fst (Id.Map.find id env.env_named_context.env_named_map) +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 = diff --git a/kernel/nativecode.ml b/kernel/nativecode.ml index 79930f6ad..bf40f7389 100644 --- a/kernel/nativecode.ml +++ b/kernel/nativecode.ml @@ -1845,7 +1845,7 @@ and compile_rel env sigma univ auxdefs n = and compile_named env sigma univ auxdefs id = let open Context.Named.Declaration in - match Context.Named.lookup id env.env_named_context.env_named_ctx with + match lookup_named id env with | LocalDef (_,t,_) -> let code = lambda_of_constr env sigma t in let auxdefs,code = compile_with_fv env sigma univ auxdefs None code in diff --git a/kernel/pre_env.ml b/kernel/pre_env.ml index 104fe59d1..7be8606ef 100644 --- a/kernel/pre_env.ml +++ b/kernel/pre_env.ml @@ -173,6 +173,9 @@ let push_named d env = indirect_pterms = env.indirect_pterms; } +let lookup_named id env = + fst (Id.Map.find id env.env_named_context.env_named_map) + let lookup_named_val id env = snd(Id.Map.find id env.env_named_context.env_named_map) diff --git a/kernel/pre_env.mli b/kernel/pre_env.mli index 1e4eff3d5..866790367 100644 --- a/kernel/pre_env.mli +++ b/kernel/pre_env.mli @@ -81,6 +81,7 @@ val map_named_val : (constr -> constr) -> named_context_val -> named_context_val val push_named : Context.Named.Declaration.t -> env -> env +val lookup_named : Id.t -> env -> Context.Named.Declaration.t val lookup_named_val : Id.t -> env -> lazy_val val env_of_named : Id.t -> env -> env diff --git a/pretyping/unification.ml b/pretyping/unification.ml index 6d80db645..36353653a 100644 --- a/pretyping/unification.ml +++ b/pretyping/unification.ml @@ -1586,7 +1586,7 @@ let make_abstraction_core name (test,out) env sigma c ty occs check_occs concl = 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 - if mem_named_context x (named_context env) then + if mem_named_context_val x (named_context_val env) then errorlabstrm "Unification.make_abstraction_core" (str "The variable " ++ Nameops.pr_id x ++ str " is already declared.") else diff --git a/proofs/logic.ml b/proofs/logic.ml index aa0b9bac6..cc95de646 100644 --- a/proofs/logic.ml +++ b/proofs/logic.ml @@ -533,7 +533,7 @@ let prim_refiner r sigma goal = nexthyp, t,cl,sigma else - (if !check && mem_named_context id (named_context_of_val sign) then + (if !check && mem_named_context_val id sign then errorlabstrm "Logic.prim_refiner" (str "Variable " ++ pr_id id ++ str " is already declared."); push_named_context_val (LocalAssum (id,t)) sign,t,cl,sigma) in diff --git a/proofs/tacmach.ml b/proofs/tacmach.ml index 50984c48e..ba22db083 100644 --- a/proofs/tacmach.ml +++ b/proofs/tacmach.ml @@ -191,9 +191,9 @@ module New = struct next_ident_away id ids let pf_get_hyp id gl = - let hyps = Proofview.Goal.hyps gl in + let hyps = Proofview.Goal.env gl in let sign = - try Context.Named.lookup id hyps + try Environ.lookup_named id hyps with Not_found -> raise (RefinerError (NoSuchHyp id)) in sign diff --git a/tactics/leminv.ml b/tactics/leminv.ml index 642bf520b..40b600c89 100644 --- a/tactics/leminv.ml +++ b/tactics/leminv.ml @@ -202,11 +202,11 @@ let inversion_scheme env sigma t sort dep_option inv_op = tclTHEN intro (onLastHypId inv_op)) pf) in let pfterm = List.hd (Proof.partial_proof pf) in - let global_named_context = Global.named_context () in + let global_named_context = Global.named_context_val () in let ownSign = ref begin fold_named_context (fun env d sign -> - if mem_named_context (get_id d) global_named_context then sign + if mem_named_context_val (get_id d) global_named_context then sign else Context.Named.add d sign) invEnv ~init:Context.Named.empty end in diff --git a/tactics/tactics.ml b/tactics/tactics.ml index f47141efb..e93a82e4d 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -179,10 +179,10 @@ let introduction ?(check=true) id = let gl = Proofview.Goal.assume gl in let concl = Proofview.Goal.concl gl in let sigma = Tacmach.New.project gl in - let hyps = Proofview.Goal.hyps gl in + let hyps = named_context_val (Proofview.Goal.env gl) in let store = Proofview.Goal.extra gl in let env = Proofview.Goal.env gl in - let () = if check && mem_named_context id hyps then + let () = if check && mem_named_context_val id hyps then errorlabstrm "Tactics.introduction" (str "Variable " ++ pr_id id ++ str " is already declared.") in @@ -507,7 +507,7 @@ let mutual_fix f n rest j = Proofview.Goal.nf_enter { enter = begin fun gl -> 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 f (named_context_of_val sign) then + if mem_named_context_val f sign then errorlabstrm "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 @@ -560,7 +560,7 @@ let mutual_cofix f others j = Proofview.Goal.nf_enter { enter = begin fun gl -> | [] -> sign | (f, ar) :: oth -> let open Context.Named.Declaration in - if mem_named_context f (named_context_of_val sign) then + 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 in @@ -2782,7 +2782,7 @@ let old_generalize_dep ?(with_let=false) c gl = let tothin = List.filter (fun id -> not (Id.List.mem id init_ids)) qhyps in let tothin' = match kind_of_term c with - | Var id when mem_named_context id sign && not (Id.List.mem id init_ids) + | Var id when mem_named_context_val id (val_of_named_context sign) && not (Id.List.mem id init_ids) -> id::tothin | _ -> tothin in @@ -2932,8 +2932,8 @@ let unfold_body x = let open Context.Named.Declaration in Proofview.Goal.enter { enter = begin fun gl -> (** We normalize the given hypothesis immediately. *) - let hyps = Proofview.Goal.hyps (Proofview.Goal.assume gl) in - let xval = match Context.Named.lookup x hyps with + let env = Proofview.Goal.env (Proofview.Goal.assume gl) in + let xval = match Environ.lookup_named x env with | LocalAssum _ -> errorlabstrm "unfold_body" (pr_id x ++ str" is not a defined hypothesis.") | LocalDef (_,xval,_) -> xval @@ -4355,7 +4355,7 @@ let induction_gen clear_flag isrec with_evars elim let cls = Option.default allHypsAndConcl cls in let t = typ_of env sigma c in let is_arg_pure_hyp = - isVar c && not (mem_named_context (destVar c) (Global.named_context())) + 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 @@ -4402,7 +4402,7 @@ let induction_gen_l isrec with_evars elim names lc = | [] -> Proofview.tclUNIT () | c::l' -> match kind_of_term c with - | Var id when not (mem_named_context id (Global.named_context())) + | Var id when not (mem_named_context_val id (Global.named_context_val ())) && not with_evars -> let _ = newlc:= id::!newlc in atomize_list l' @@ -4824,7 +4824,7 @@ let abstract_subproof id gk tac = let open Context.Named.Declaration in Proofview.Goal.nf_s_enter { s_enter = begin fun gl -> let sigma = Proofview.Goal.sigma gl in - let current_sign = Global.named_context() + let current_sign = Global.named_context_val () and global_sign = Proofview.Goal.hyps gl in let sigma = Sigma.to_evar_map sigma in let evdref = ref sigma in @@ -4832,8 +4832,8 @@ let abstract_subproof id gk tac = List.fold_right (fun d (s1,s2) -> let id = get_id d in - if mem_named_context id current_sign && - interpretable_as_section_decl evdref (Context.Named.lookup id current_sign) d + if mem_named_context_val id current_sign && + interpretable_as_section_decl evdref (lookup_named_val id current_sign) d then (s1,push_named_context_val d s2) else (Context.Named.add d s1,s2)) global_sign (Context.Named.empty, empty_named_context_val) in -- cgit v1.2.3